all: remove theory import path references

In Isabelle2020, when isabelle jedit is started without a session
context, e.g. `isabelle jedit -l ASpec`, theory imports with path
references cause the isabelle process to hang.

Since sessions now declare directories, Isabelle can find those files
without path reference and we therefore remove all such path references
from import statements. With this, `jedit` and `build` should work with
and without explicit session context as before.

Signed-off-by: Gerwin Klein <gerwin.klein@data61.csiro.au>
This commit is contained in:
Gerwin Klein 2020-10-31 16:30:58 +10:00 committed by Gerwin Klein
parent b5e7fa4e45
commit a45adef66a
453 changed files with 596 additions and 598 deletions

View File

@ -7,7 +7,7 @@ chapter \<open>\label{h:dataportbase}Example -- Dataports\<close>
(*<*)
(* THIS FILE IS AUTOMATICALLY GENERATED. YOUR EDITS WILL BE OVERWRITTEN. *)
theory GenDataportBase
imports "../Types" "../Abbreviations" "../Connector"
imports "Types" "Abbreviations" "Connector"
begin
(*>*)

View File

@ -5,7 +5,7 @@
*)
(*<*)
theory UserDataport
imports "../UserStubs" GenDataportBase
imports "UserStubs" GenDataportBase
begin
(* Stub to expose existing static UserStubs.thy. This is to avoid having to

View File

@ -7,7 +7,7 @@ chapter \<open>\label{s:eventbase}Example -- Events\<close>
(*<*)
(* THIS FILE IS AUTOMATICALLY GENERATED. YOUR EDITS WILL BE OVERWRITTEN. *)
theory GenEventBase
imports "../Types" "../Abbreviations" "../Connector"
imports "Types" "Abbreviations" "Connector"
begin
(*>*)

View File

@ -5,7 +5,7 @@
*)
(*<*)
theory UserEvent
imports "../UserStubs" GenEventBase
imports "UserStubs" GenEventBase
begin
(* Stub to expose existing static UserStubs.thy. This is to avoid having to

View File

@ -7,7 +7,7 @@ chapter \<open>\label{h:procbase}Example -- Procedures\<close>
(*<*)
(* THIS FILE IS AUTOMATICALLY GENERATED. YOUR EDITS WILL BE OVERWRITTEN. *)
theory GenSimpleBase
imports "../Types" "../Abbreviations" "../Connector"
imports "Types" "Abbreviations" "Connector"
begin
(*>*)

View File

@ -5,7 +5,7 @@
*)
(*<*)
theory UserSimple
imports "../UserStubs" GenSimpleBase
imports "UserStubs" GenSimpleBase
begin
(* Stub to expose existing static UserStubs.thy. This is to avoid having to

View File

@ -6,7 +6,7 @@
(*<*)
(* THIS FILE IS AUTOMATICALLY GENERATED. YOUR EDITS WILL BE OVERWRITTEN. *)
theory GenFilter2Base
imports "../Types" "../Abbreviations" "../Connector"
imports "Types" "Abbreviations" "Connector"
begin
(* Connections *)

View File

@ -5,7 +5,7 @@
*)
(*<*)
theory UserFilter2
imports GenFilter2Base "../UserStubs"
imports GenFilter2Base "UserStubs"
begin
(*>*)

View File

@ -8,7 +8,7 @@ chapter \<open>\label{h:filter}Example -- System Level Reasoning\<close>
(*<*)
(* THIS FILE IS AUTOMATICALLY GENERATED. YOUR EDITS WILL BE OVERWRITTEN. *)
theory GenFilterBase
imports "../Types" "../Abbreviations" "../Connector"
imports "Types" "Abbreviations" "Connector"
begin
(*>*)

View File

@ -5,7 +5,7 @@
*)
(*<*)
theory UserFilter
imports "../UserStubs" GenFilterBase
imports "UserStubs" GenFilterBase
begin
(* Stub to expose existing static UserStubs.thy. This is to avoid having to

View File

@ -9,7 +9,7 @@
theory Apply_Trace
imports
Main
"ml-helpers/MLUtils"
"MLUtils"
begin

View File

@ -5,7 +5,7 @@
*)
theory AutoLevity (* FIXME: broken *)
imports "../tools/proofcount/ProofGraph"
imports "ProofGraph"
begin
ML \<open>val infoflow_file = "~~/../lib/proof_counting/Noninterference_Refinement_graphs.xml"\<close>
ML \<open>val (full_spec,proof_spec_raw,thy_deps) = Proof_Graph.read_graph_spec_from infoflow_file\<close>

View File

@ -6,7 +6,7 @@
theory BCorres_UL
imports
"Monad_WP/NonDetMonadVCG"
NonDetMonadVCG
Crunch_Instances_NonDet
begin

View File

@ -8,7 +8,7 @@
theory Bisim_UL
imports
"Monad_WP/NonDetMonadVCG"
NonDetMonadVCG
Corres_UL
EmptyFailLib
begin

View File

@ -7,8 +7,8 @@
theory Corres_UL
imports
Crunch_Instances_NonDet
"Monad_WP/wp/WPEx"
"Monad_WP/wp/WPFix"
WPEx
WPFix
HaskellLemmaBucket
begin

View File

@ -6,9 +6,9 @@
theory Crunch
imports
"Monad_WP/wp/WPSimp"
WPSimp
Lib
"ml-helpers/MLUtils"
MLUtils
keywords "crunch" "crunch_ignore" "crunches" :: thy_decl
begin

View File

@ -7,8 +7,8 @@
theory Crunch_Instances_NonDet
imports
Crunch
"Monad_WP/wp/WPEx"
"Monad_WP/NonDetMonadVCG"
WPEx
NonDetMonadVCG
begin
lemmas [crunch_param_rules] = Let_def return_bind returnOk_bindE

View File

@ -7,7 +7,7 @@
theory Crunch_Instances_Trace
imports
Crunch
"Monad_WP/TraceMonadVCG"
TraceMonadVCG
begin
lemmas [crunch_param_rules] = Let_def return_bind returnOk_bindE

View File

@ -10,7 +10,7 @@
theory Eisbach_Methods
imports
"subgoal_focus/Subgoal_Methods"
Subgoal_Methods
"HOL-Eisbach.Eisbach_Tools"
Rule_By_Method
Local_Method

View File

@ -6,7 +6,7 @@
theory EmptyFailLib
imports
"Monad_WP/NonDetMonad"
NonDetMonad
HaskellLib_H
begin

View File

@ -7,7 +7,7 @@
theory FP_Eval
imports
HOL.HOL
"ml-helpers/TermPatternAntiquote"
TermPatternAntiquote
begin
text \<open>

View File

@ -7,7 +7,7 @@
theory GenericLib
imports
Crunch_Instances_NonDet
"Monad_WP/wp/WPEx"
WPEx
HaskellLemmaBucket
begin

View File

@ -13,8 +13,8 @@ theory HaskellLib_H
imports
Lib
NatBitwise
"More_Numeral_Type"
"Monad_WP/NonDetMonadVCG"
More_Numeral_Type
NonDetMonadVCG
begin
abbreviation (input) "flip \<equiv> swp"

View File

@ -7,7 +7,7 @@
theory ML_Goal_Test
imports
ML_Goal
"ml-helpers/MLUtils"
MLUtils
begin
experiment begin

View File

@ -20,7 +20,7 @@
* are added to the "monad_eq" set.
*)
theory MonadEq
imports "Monad_WP/NonDetMonadVCG"
imports NonDetMonadVCG
begin
(* Setup "monad_eq" attributes. *)

View File

@ -6,8 +6,8 @@
theory Datatype_Schematic
imports
"../ml-helpers/MLUtils"
"../ml-helpers/TermPatternAntiquote"
MLUtils
TermPatternAntiquote
begin
text \<open>

View File

@ -11,7 +11,7 @@
chapter "Nondeterministic State Monad with Failure"
theory NonDetMonad
imports "../Lib"
imports Lib
begin
text \<open>

View File

@ -7,8 +7,8 @@
theory NonDetMonadVCG
imports
NonDetMonadLemmas
"wp/WPSimp"
"Strengthen"
WPSimp
Strengthen
begin
declare K_def [simp]

View File

@ -12,8 +12,8 @@
theory OptionMonad (* FIXME: this is really a Reader_Option_Monad *)
imports
"../Lib" (* FIXME: reduce dependencies *)
"Less_Monad_Syntax"
Lib (* FIXME: reduce dependencies *)
Less_Monad_Syntax
begin
type_synonym ('s,'a) lookup = "'s \<Rightarrow> 'a option"

View File

@ -13,7 +13,7 @@ This list is almost certainly incomplete; add rules here as they are needed.
theory OptionMonadWP
imports
OptionMonadND
"wp/WP"
WP
begin
declare K_def [simp]

View File

@ -5,8 +5,8 @@
*)
theory TraceMonad
imports
"../Lib"
"Strengthen"
Lib
Strengthen
begin
text \<open>

View File

@ -6,8 +6,8 @@
theory TraceMonadVCG
imports
TraceMonad
"wp/WPSimp"
"Strengthen"
WPSimp
Strengthen
begin
lemma trace_steps_append:

View File

@ -8,11 +8,11 @@
theory Eisbach_WP
imports
"../../Eisbach_Methods"
"../NonDetMonadVCG"
"../../Conjuncts"
"../../Rule_By_Method"
"WPI"
Eisbach_Methods
NonDetMonadVCG
Conjuncts
Rule_By_Method
WPI
begin

View File

@ -6,10 +6,10 @@
theory WP
imports
"WP_Pre"
"WPFix"
"../../Apply_Debug"
"../../ml-helpers/MLUtils"
WP_Pre
WPFix
Apply_Debug
MLUtils
begin
definition

View File

@ -7,8 +7,8 @@
theory WPBang
imports
WP
"../../ProvePart"
"../NonDetMonadVCG"
ProvePart
NonDetMonadVCG
begin
lemma conj_meta_forward:

View File

@ -6,8 +6,8 @@
theory WPEx
imports
"../NonDetMonadVCG"
"../Strengthen"
NonDetMonadVCG
Strengthen
begin
text \<open>WPEx - the WP Extension Experiment\<close>

View File

@ -6,8 +6,8 @@
theory WPFix
imports
"../Datatype_Schematic"
"../Strengthen"
Datatype_Schematic
Strengthen
begin

View File

@ -32,9 +32,9 @@
theory WPI
imports
"../../Eisbach_Methods"
"../NonDetMonadLemmas"
"WPEx"
Eisbach_Methods
NonDetMonadLemmas
WPEx
begin
text \<open>The ML version of repeat_new is slightly faster than the Eisbach one.\<close>

View File

@ -9,7 +9,7 @@ imports
"WP"
"WPC"
"WPFix"
"../../Simp_No_Conditional"
Simp_No_Conditional
begin
(* Wrap up the standard usage pattern of wp/wpc/simp into its own command: *)

View File

@ -7,7 +7,7 @@
theory WP_Pre
imports
Main
"../../Trace_Schematic_Insts"
Trace_Schematic_Insts
"HOL-Eisbach.Eisbach_Tools"
begin

View File

@ -8,7 +8,7 @@
theory MonadicRewrite
imports
"Monad_WP/NonDetMonadVCG"
NonDetMonadVCG
Corres_UL
EmptyFailLib
LemmaBucket

View File

@ -6,9 +6,9 @@
theory NonDetMonadLemmaBucket
imports
"Monad_WP/NonDetMonadVCG"
"MonadEq"
"Monad_WP/WhileLoopRulesCompleteness"
NonDetMonadVCG
MonadEq
WhileLoopRulesCompleteness
"Word_Lib.Distinct_Prop"
begin
setup \<open>AutoLevity_Base.add_attribute_test "wp" WeakestPre.is_wp_rule\<close>

View File

@ -8,7 +8,7 @@ theory RangeMap
imports
FastMap
FP_Eval
"ml-helpers/MkTermAntiquote"
MkTermAntiquote
begin
text \<open>

View File

@ -7,8 +7,8 @@
theory Trace_Schematic_Insts
imports
Main
"ml-helpers/MLUtils"
"ml-helpers/TermPatternAntiquote"
MLUtils
TermPatternAntiquote
begin
text \<open>

View File

@ -7,8 +7,8 @@
theory WordSetup (* part of non-AFP Word_Lib *)
imports
"../Distinct_Prop"
"../Word_Lemmas_32_Internal"
Distinct_Prop
Word_Lemmas_32_Internal
begin
(* Distinct_Prop lemmas that need word lemmas: *)

View File

@ -7,8 +7,8 @@
theory WordSetup (* part of non-AFP Word_Lib *)
imports
"../Distinct_Prop"
"../Word_Lemmas_32_Internal"
Distinct_Prop
Word_Lemmas_32_Internal
begin
(* Distinct_Prop lemmas that need word lemmas: *)

View File

@ -7,8 +7,8 @@
theory WordSetup (* part of non-AFP Word_Lib *)
imports
"../Distinct_Prop"
"../Word_Lemmas_64_Internal"
Distinct_Prop
Word_Lemmas_64_Internal
begin
(* Distinct_Prop lemmas that need word lemmas: *)

View File

@ -7,8 +7,8 @@
theory WordSetup (* part of non-AFP Word_Lib *)
imports
"../Distinct_Prop"
"../Word_Lemmas_64_Internal"
Distinct_Prop
Word_Lemmas_64_Internal
begin
(* Distinct_Prop lemmas that need word lemmas: *)

View File

@ -5,7 +5,7 @@
*)
theory Peterson_Atomicity
imports
"../Atomicity_Lib"
Atomicity_Lib
begin
text \<open>

View File

@ -5,7 +5,7 @@
*)
theory Plus2_Prefix
imports
"../Prefix_Refinement"
Prefix_Refinement
begin
section \<open>The plus2 example.\<close>

View File

@ -10,7 +10,7 @@
*)
theory Sep_Tactics_Test
imports "../Sep_Tactics"
imports Sep_Tactics
begin
text \<open>Substitution and forward/backward reasoning\<close>

View File

@ -15,8 +15,8 @@ chapter "Example from HOL/Hoare/Separation"
theory Simple_Separation_Example
imports
"HOL-Hoare.Hoare_Logic_Abort"
"../Sep_Heap_Instance"
"../Sep_Tactics"
Sep_Heap_Instance
Sep_Tactics
begin
declare [[syntax_ambiguity_warning = false]]

View File

@ -14,8 +14,8 @@ chapter "Separation Algebra for Virtual Memory"
theory VM_Example
imports
"../Sep_Tactics"
"../Map_Extra"
Sep_Tactics
Map_Extra
begin
text \<open>

View File

@ -13,9 +13,9 @@ chapter "Instantiating capDL as a separation algebra."
theory Abstract_Separation_D
imports
"../../Sep_Tactics"
Sep_Tactics
Types_D
"../../Map_Extra"
Map_Extra
begin
(**************************************

View File

@ -8,7 +8,7 @@
(* Arch specific lemmas that should be moved into theory files before CRefine *)
theory ArchMove_C
imports "../Move_C"
imports Move_C
begin
lemma ps_clear_is_aligned_ksPSpace_None:

View File

@ -8,7 +8,7 @@
theory Ctac_lemmas_C
imports
"../lib/Ctac"
Ctac
begin
context kernel

View File

@ -7,7 +7,7 @@
chapter "Toplevel Refinement Statement"
theory Refine_C
imports Init_C Fastpath_C "../lib/CToCRefine"
imports Init_C Fastpath_C CToCRefine
begin
context begin interpretation Arch . (*FIXME: arch_split*)

View File

@ -8,7 +8,7 @@
(* Arch specific lemmas that should be moved into theory files before CRefine *)
theory ArchMove_C
imports "../Move_C"
imports Move_C
begin
(* FIXME move: need a theory on top of CSpec that arches can share *)

View File

@ -8,7 +8,7 @@
theory Ctac_lemmas_C
imports
"../lib/Ctac"
Ctac
begin
context kernel

View File

@ -7,7 +7,7 @@
chapter "Toplevel Refinement Statement"
theory Refine_C
imports Init_C Fastpath_C "../lib/CToCRefine"
imports Init_C Fastpath_C CToCRefine
begin
context begin interpretation Arch . (*FIXME: arch_split*)

View File

@ -7,7 +7,7 @@
(* Arch specific lemmas that should be moved into theory files before CRefine *)
theory ArchMove_C
imports "../Move_C"
imports Move_C
begin

View File

@ -8,7 +8,7 @@
theory Ctac_lemmas_C
imports
"../lib/Ctac"
Ctac
begin
context kernel

View File

@ -17,7 +17,7 @@ imports
"Refine.RAB_FN"
"CLib.MonadicRewrite_C"
"../lib/CToCRefine"
CToCRefine
begin
context begin interpretation Arch . (*FIXME: arch_split*)

View File

@ -8,7 +8,7 @@
(* Arch specific lemmas that should be moved into theory files before CRefine *)
theory ArchMove_C
imports "../Move_C"
imports Move_C
begin
lemma ps_clear_is_aligned_ksPSpace_None:

View File

@ -8,7 +8,7 @@
theory Ctac_lemmas_C
imports
"../lib/Ctac"
Ctac
begin
context kernel

View File

@ -17,7 +17,7 @@ imports
"Refine.RAB_FN"
"CLib.MonadicRewrite_C"
"../lib/CToCRefine"
CToCRefine
begin
context begin interpretation Arch . (*FIXME: arch_split*)

View File

@ -7,7 +7,7 @@
theory Corres_C
imports
"CLib.CCorresLemmas"
"../$L4V_ARCH/SR_lemmas_C"
SR_lemmas_C
begin
abbreviation

View File

@ -8,7 +8,7 @@ chapter \<open>Abstract datatype for the abstract specification\<close>
theory ADT_AI
imports
"./$L4V_ARCH/ArchADT_AI"
ArchADT_AI
begin
context begin interpretation Arch .

View File

@ -9,7 +9,7 @@ The main theorem
*)
theory AInvs
imports "./$L4V_ARCH/ArchAInvsPre"
imports ArchAInvsPre
begin
lemma st_tcb_at_nostate_upd:

View File

@ -5,7 +5,7 @@
*)
theory AInvsPre
imports "./$L4V_ARCH/ArchVSpaceEntries_AI" ADT_AI
imports ArchVSpaceEntries_AI ADT_AI
begin
locale AInvsPre =

View File

@ -9,7 +9,7 @@ chapter \<open>ARM-specific definitions for abstract datatype for the abstract s
theory ArchADT_AI
imports
"Lib.Simulation"
"../Invariants_AI"
Invariants_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchAInvsPre
imports "../AInvsPre"
imports AInvsPre
begin
context Arch begin

View File

@ -9,8 +9,9 @@ Lemmas on arch get/set object etc
*)
theory ArchAcc_AI
imports "../SubMonad_AI"
"Lib.Crunch_Instances_NonDet"
imports
SubMonad_AI
"Lib.Crunch_Instances_NonDet"
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchArch_AI
imports "../Arch_AI"
imports Arch_AI
begin
context Arch begin global_naming ARM

View File

@ -6,7 +6,7 @@
theory ArchBCorres2_AI
imports
"../BCorres2_AI"
BCorres2_AI
begin
context Arch begin global_naming ARM

View File

@ -6,7 +6,7 @@
theory ArchBCorres_AI
imports
"../BCorres_AI"
BCorres_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchBits_AI
imports "../Invariants_AI"
imports Invariants_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchCNodeInv_AI
imports "../CNodeInv_AI"
imports CNodeInv_AI
begin
context Arch begin global_naming ARM

View File

@ -9,7 +9,7 @@ CSpace invariants
*)
theory ArchCSpaceInvPre_AI
imports "../CSpaceInvPre_AI"
imports CSpaceInvPre_AI
begin
context Arch begin global_naming ARM

View File

@ -9,7 +9,7 @@ CSpace invariants
*)
theory ArchCSpaceInv_AI
imports "../CSpaceInv_AI"
imports CSpaceInv_AI
begin
context Arch begin global_naming ARM

View File

@ -9,7 +9,7 @@ ARM-specific CSpace invariants
*)
theory ArchCSpacePre_AI
imports "../CSpacePre_AI"
imports CSpacePre_AI
begin
context Arch begin global_naming ARM

View File

@ -9,7 +9,7 @@ ARM-specific CSpace invariants
*)
theory ArchCSpace_AI
imports "../CSpace_AI"
imports CSpace_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchDetSchedAux_AI
imports "../DetSchedAux_AI"
imports DetSchedAux_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchDetSchedDomainTime_AI
imports "../DetSchedDomainTime_AI"
imports DetSchedDomainTime_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchDetSchedSchedule_AI
imports "../DetSchedSchedule_AI"
imports DetSchedSchedule_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchDeterministic_AI
imports "../Deterministic_AI"
imports Deterministic_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchDetype_AI
imports "../Detype_AI"
imports Detype_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchEmptyFail_AI
imports "../EmptyFail_AI"
imports EmptyFail_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchFinalise_AI
imports "../Finalise_AI"
imports Finalise_AI
begin
context Arch begin

View File

@ -9,7 +9,7 @@
*)
theory ArchInterruptAcc_AI
imports "../InterruptAcc_AI"
imports InterruptAcc_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchInterrupt_AI
imports "../Interrupt_AI"
imports Interrupt_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchInvariants_AI
imports "../InvariantsPre_AI"
imports InvariantsPre_AI
begin
\<comment> \<open>---------------------------------------------------------------------------\<close>

View File

@ -5,7 +5,7 @@
*)
theory ArchIpcCancel_AI
imports "../IpcCancel_AI"
imports IpcCancel_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchIpc_AI
imports "../Ipc_AI"
imports Ipc_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchKHeap_AI
imports "../KHeapPre_AI"
imports KHeapPre_AI
begin
context Arch begin global_naming ARM

View File

@ -9,9 +9,9 @@
theory ArchKernelInit_AI
imports
"../ADT_AI"
"../Tcb_AI"
"../Arch_AI"
ADT_AI
Tcb_AI
Arch_AI
begin
context Arch begin global_naming ARM (*FIXME: arch_split*)

View File

@ -9,7 +9,7 @@
*)
theory ArchRetype_AI
imports "../Retype_AI"
imports Retype_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchSchedule_AI
imports "../Schedule_AI"
imports Schedule_AI
begin
context Arch begin global_naming ARM

View File

@ -10,7 +10,7 @@ Refinement for handleEvent and syscalls
theory ArchSyscall_AI
imports
"../Syscall_AI"
Syscall_AI
begin
context Arch begin global_naming ARM

View File

@ -5,7 +5,7 @@
*)
theory ArchTcbAcc_AI
imports "../TcbAcc_AI"
imports TcbAcc_AI
begin
context Arch begin global_naming ARM

Some files were not shown because too many files have changed in this diff Show More