lib: theory import fixes for new sessions

Signed-off-by: Gerwin Klein <gerwin.klein@proofcraft.systems>
This commit is contained in:
Gerwin Klein 2023-01-23 11:43:26 +11:00
parent 0aa42207e5
commit b92974d93f
No known key found for this signature in database
GPG Key ID: 20A847CE6AB7F5F3
22 changed files with 32 additions and 36 deletions

View File

@ -7,8 +7,8 @@
theory Corres_UL theory Corres_UL
imports imports
Crunch_Instances_NonDet Crunch_Instances_NonDet
WPEx Monads.WPEx
WPFix Monads.WPFix
HaskellLemmaBucket HaskellLemmaBucket
begin begin

View File

@ -6,7 +6,7 @@
theory Crunch theory Crunch
imports imports
WPSimp Monads.WPSimp
Lib Lib
ML_Utils.ML_Utils ML_Utils.ML_Utils
keywords "crunch" "crunch_ignore" "crunches" :: thy_decl keywords "crunch" "crunch_ignore" "crunches" :: thy_decl

View File

@ -7,7 +7,7 @@
theory Crunch_Instances_Trace theory Crunch_Instances_Trace
imports imports
Crunch Crunch
TraceMonadVCG Monads.TraceMonadVCG
begin begin
lemmas [crunch_param_rules] = Let_def return_bind returnOk_bindE lemmas [crunch_param_rules] = Let_def return_bind returnOk_bindE

View File

@ -55,7 +55,7 @@ In this sense, EquivValid statements could be thought of as \<^emph>\<open>relat
text \<open> This tutorial will introduce some syntactic sugar to emphasise the similarity text \<open> This tutorial will introduce some syntactic sugar to emphasise the similarity
between Hoare triples and EquivValid statements. between Hoare triples and EquivValid statements.
(Here, \<open>\<top>\<top>\<close> is an alias provided by Lib.NonDetMonad for the trivial binary predicate, (Here, \<open>\<top>\<top>\<close> is an alias provided by Monads.NonDetMonad for the trivial binary predicate,
which always returns \<open>True\<close>; similarly, there is also \<open>\<bottom>\<bottom>\<close> for \<open>False\<close>.) \<close> which always returns \<open>True\<close>; similarly, there is also \<open>\<bottom>\<bottom>\<close> for \<open>False\<close>.) \<close>
abbreviation abbreviation
@ -140,7 +140,7 @@ where
text \<open> We could just as well define equivalences similarly for text \<open> We could just as well define equivalences similarly for
\<open>inbox\<close>, \<open>outbox\<close>, or \<open>job\<close> -- we omit these here. \<close> \<open>inbox\<close>, \<open>outbox\<close>, or \<open>job\<close> -- we omit these here. \<close>
text \<open> Furthermore, you can use the \<open>and\<close> alias provided by @{theory Lib.Fun_Pred_Syntax} text \<open> Furthermore, you can use the \<open>and\<close> alias provided by @{theory Monads.Fun_Pred_Syntax}
to write the conjunction of two binary predicates. to write the conjunction of two binary predicates.
(Similarly, there is also an \<open>or\<close> alias for disjunction.)\<close> (Similarly, there is also an \<open>or\<close> alias for disjunction.)\<close>

View File

@ -6,7 +6,7 @@
theory EmptyFailLib theory EmptyFailLib
imports imports
NonDetMonad Monads.NonDetMonad
HaskellLib_H HaskellLib_H
begin begin

View File

@ -7,7 +7,7 @@
theory Extract_Conjunct theory Extract_Conjunct
imports imports
"Main" "Main"
"Eisbach_Methods" Eisbach_Tools.Eisbach_Methods
begin begin
section \<open>Extracting conjuncts in the conclusion\<close> section \<open>Extracting conjuncts in the conclusion\<close>

View File

@ -7,7 +7,7 @@
theory GenericLib theory GenericLib
imports imports
Crunch_Instances_NonDet Crunch_Instances_NonDet
WPEx Monads.WPEx
HaskellLemmaBucket HaskellLemmaBucket
begin begin

View File

@ -6,8 +6,8 @@
theory Guess_ExI theory Guess_ExI
imports imports
Eisbach_Methods Eisbach_Tools.Eisbach_Methods
Apply_Debug Eisbach_Tools.Apply_Debug
begin begin
(* (*

View File

@ -6,7 +6,7 @@
theory LexordList imports theory LexordList imports
Main Main
Eisbach_Methods (* for tests *) Eisbach_Tools.Eisbach_Methods (* for tests *)
begin begin
text \<open> text \<open>

View File

@ -18,8 +18,8 @@ imports
Extract_Conjunct Extract_Conjunct
ML_Goal ML_Goal
Eval_Bool Eval_Bool
Fun_Pred_Syntax Monads.Fun_Pred_Syntax
Monad_Lib Monads.Monad_Lib
NICTATools NICTATools
"Word_Lib.WordSetup" "Word_Lib.WordSetup"
begin begin

View File

@ -9,7 +9,7 @@
Should not be Hoare triples (those go into a different theory). *) Should not be Hoare triples (those go into a different theory). *)
theory Monad_Equations theory Monad_Equations
imports MonadEq_Lemmas imports Monads.MonadEq_Lemmas
begin begin
lemmas assertE_assert = assertE_liftE lemmas assertE_assert = assertE_liftE

View File

@ -7,16 +7,16 @@
(* Miscellaneous Isabelle tools. *) (* Miscellaneous Isabelle tools. *)
theory NICTATools theory NICTATools
imports imports
Apply_Trace_Cmd Eisbach_Tools.Apply_Trace_Cmd
Apply_Debug Eisbach_Tools.Apply_Debug
Find_Names Find_Names
(* Solves_Tac *) (* Solves_Tac *)
Rule_By_Method Eisbach_Tools.Rule_By_Method
Eisbach_Methods Eisbach_Tools.Eisbach_Methods
Time_Methods_Cmd Time_Methods_Cmd
Try_Attribute Try_Attribute
Repeat_Attribute Repeat_Attribute
Trace_Schematic_Insts Eisbach_Tools.Trace_Schematic_Insts
Insulin Insulin
ShowTypes ShowTypes
Locale_Abbrev Locale_Abbrev

View File

@ -19,11 +19,9 @@ session Lib (lib) = Word_Lib +
List_Lib List_Lib
SubMonadLib SubMonadLib
Simulation Simulation
MonadEq
SimpStrategy SimpStrategy
Extract_Conjunct Extract_Conjunct
GenericLib GenericLib
ProvePart
Corres_Adjust_Preconds Corres_Adjust_Preconds
Requalify Requalify
Value_Abbreviation Value_Abbreviation
@ -41,7 +39,6 @@ session Lib (lib) = Word_Lib +
Try_Methods Try_Methods
ListLibLemmas ListLibLemmas
Time_Methods_Cmd Time_Methods_Cmd
Apply_Debug
MonadicRewrite MonadicRewrite
HaskellLemmaBucket HaskellLemmaBucket
FP_Eval FP_Eval
@ -122,7 +119,6 @@ session LibTest (lib) in test = Refine +
RangeMap_Test RangeMap_Test
FP_Eval_Tests FP_Eval_Tests
Trace_Schematic_Insts_Test Trace_Schematic_Insts_Test
Local_Method_Tests
Qualify_Test Qualify_Test
Locale_Abbrev_Test Locale_Abbrev_Test
Value_Type_Test Value_Type_Test

View File

@ -6,7 +6,7 @@
theory Try_Methods theory Try_Methods
imports Eisbach_Methods imports Eisbach_Tools.Eisbach_Methods
keywords "trym" :: diag keywords "trym" :: diag
and "add_try_method" :: thy_decl and "add_try_method" :: thy_decl

View File

@ -6,7 +6,7 @@
theory BitFieldProofsLib theory BitFieldProofsLib
imports imports
"Lib.Eisbach_Methods" "Eisbach_Tools.Eisbach_Methods"
TypHeapLib TypHeapLib
begin begin

View File

@ -10,8 +10,8 @@ text \<open>A simple proof method for rewriting Simpl programs under a predicate
theory Simpl_Rewrite theory Simpl_Rewrite
imports imports
"Simpl-VCG.Vcg" "Simpl-VCG.Vcg"
"Lib.Eisbach_Methods" "Eisbach_Tools.Eisbach_Methods"
"Lib.Apply_Debug" "Eisbach_Tools.Apply_Debug"
begin begin
text \<open>One layer of context around a Simpl program. text \<open>One layer of context around a Simpl program.

View File

@ -7,7 +7,7 @@ theory Prefix_Refinement
imports imports
Triv_Refinement Triv_Refinement
"Lib.TraceMonadLemmas" "Monads.TraceMonadLemmas"
begin begin

View File

@ -6,8 +6,8 @@
theory Triv_Refinement theory Triv_Refinement
imports imports
"Lib.TraceMonadVCG" "Monads.TraceMonadVCG"
"Lib.Strengthen" "Monads.Strengthen"
begin begin

View File

@ -6,8 +6,8 @@
theory Apply_Debug_Test theory Apply_Debug_Test
imports imports
Lib.Apply_Debug Eisbach_Tools.Apply_Debug
Lib.Apply_Trace_Cmd Eisbach_Tools.Apply_Trace_Cmd
begin begin
chapter \<open>Apply_Debug\<close> chapter \<open>Apply_Debug\<close>

View File

@ -7,7 +7,7 @@ theory Match_Abbreviation_Test
imports imports
Lib.Match_Abbreviation Lib.Match_Abbreviation
Lib.NonDetMonad Monads.NonDetMonad
begin begin
experiment experiment

View File

@ -6,7 +6,7 @@
theory Time_Methods_Cmd_Test imports theory Time_Methods_Cmd_Test imports
Lib.Time_Methods_Cmd Lib.Time_Methods_Cmd
Lib.Eisbach_Methods Eisbach_Tools.Eisbach_Methods
"HOL-Library.Sublist" "HOL-Library.Sublist"
begin begin

View File

@ -6,7 +6,7 @@
theory Trace_Schematic_Insts_Test theory Trace_Schematic_Insts_Test
imports imports
Lib.Trace_Schematic_Insts Eisbach_Tools.Trace_Schematic_Insts
begin begin
text \<open> text \<open>