| author | paulson <lp15@cam.ac.uk> | 
| Wed, 07 Feb 2024 11:52:34 +0000 | |
| changeset 79583 | a521c241e946 | 
| parent 74561 | 8e6c973003c8 | 
| child 82581 | 0133fe6a1f55 | 
| permissions | -rw-r--r-- | 
| 6778 | 1  | 
(* Title: Pure/Isar/calculation.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
4  | 
Generic calculational proofs.  | 
| 6778 | 5  | 
*)  | 
6  | 
||
7  | 
signature CALCULATION =  | 
|
8  | 
sig  | 
|
| 21506 | 9  | 
val print_rules: Proof.context -> unit  | 
| 63543 | 10  | 
val check: Proof.state -> thm list option  | 
| 18728 | 11  | 
val trans_add: attribute  | 
12  | 
val trans_del: attribute  | 
|
13  | 
val sym_add: attribute  | 
|
14  | 
val sym_del: attribute  | 
|
15  | 
val symmetric: attribute  | 
|
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
16  | 
val also: thm list option -> bool -> Proof.state -> Proof.state Seq.result Seq.seq  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
56894 
diff
changeset
 | 
17  | 
val also_cmd: (Facts.ref * Token.src list) list option ->  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
18  | 
bool -> Proof.state -> Proof.state Seq.result Seq.seq  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
19  | 
val finally: thm list option -> bool -> Proof.state -> Proof.state Seq.result Seq.seq  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
56894 
diff
changeset
 | 
20  | 
val finally_cmd: (Facts.ref * Token.src list) list option -> bool ->  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
21  | 
Proof.state -> Proof.state Seq.result Seq.seq  | 
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
22  | 
val moreover: bool -> Proof.state -> Proof.state  | 
| 
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
23  | 
val ultimately: bool -> Proof.state -> Proof.state  | 
| 6778 | 24  | 
end;  | 
25  | 
||
26  | 
structure Calculation: CALCULATION =  | 
|
27  | 
struct  | 
|
28  | 
||
| 18637 | 29  | 
(** calculation data **)  | 
| 6778 | 30  | 
|
| 63543 | 31  | 
type calculation = {result: thm list, level: int, serial: serial, pos: Position.T};
 | 
32  | 
||
| 
37047
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
33  | 
structure Data = Generic_Data  | 
| 18637 | 34  | 
(  | 
| 74149 | 35  | 
type T = (thm Item_Net.T * thm Item_Net.T) * calculation option;  | 
| 74152 | 36  | 
val empty = ((Thm.item_net_elim, Thm.item_net), NONE);  | 
| 33519 | 37  | 
fun merge (((trans1, sym1), _), ((trans2, sym2), _)) =  | 
| 74149 | 38  | 
((Item_Net.merge (trans1, trans2), Item_Net.merge (sym1, sym2)), NONE);  | 
| 18637 | 39  | 
);  | 
40  | 
||
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
41  | 
val get_rules = #1 o Data.get o Context.Proof;  | 
| 74149 | 42  | 
val get_trans_rules = Item_Net.content o #1 o get_rules;  | 
43  | 
val get_sym_rules = Item_Net.content o #2 o get_rules;  | 
|
| 63543 | 44  | 
val get_calculation = #2 o Data.get o Context.Proof;  | 
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
45  | 
|
| 22846 | 46  | 
fun print_rules ctxt =  | 
| 74149 | 47  | 
let val pretty_thm = Thm.pretty_thm_item ctxt in  | 
48  | 
[Pretty.big_list "transitivity rules:" (map pretty_thm (get_trans_rules ctxt)),  | 
|
49  | 
Pretty.big_list "symmetry rules:" (map pretty_thm (get_sym_rules ctxt))]  | 
|
| 
56334
 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 
wenzelm 
parents: 
55141 
diff
changeset
 | 
50  | 
end |> Pretty.writeln_chunks;  | 
| 6778 | 51  | 
|
52  | 
||
53  | 
(* access calculation *)  | 
|
54  | 
||
| 63543 | 55  | 
fun check_calculation state =  | 
56  | 
(case get_calculation (Proof.context_of state) of  | 
|
| 15531 | 57  | 
NONE => NONE  | 
| 63543 | 58  | 
| SOME calculation =>  | 
59  | 
if #level calculation = Proof.level state then SOME calculation else NONE);  | 
|
60  | 
||
61  | 
val check = Option.map #result o check_calculation;  | 
|
| 6778 | 62  | 
|
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
63  | 
val calculationN = "calculation";  | 
| 6778 | 64  | 
|
| 63543 | 65  | 
fun update_calculation calc state =  | 
66  | 
let  | 
|
67  | 
fun report def serial pos =  | 
|
68  | 
Context_Position.report (Proof.context_of state)  | 
|
69  | 
(Position.thread_data ())  | 
|
| 74183 | 70  | 
        (Position.make_entity_markup def serial calculationN ("", pos));
 | 
| 63543 | 71  | 
val calculation =  | 
72  | 
(case calc of  | 
|
73  | 
NONE => NONE  | 
|
74  | 
| SOME result =>  | 
|
75  | 
(case check_calculation state of  | 
|
76  | 
NONE =>  | 
|
77  | 
let  | 
|
78  | 
val level = Proof.level state;  | 
|
79  | 
val serial = serial ();  | 
|
80  | 
val pos = Position.thread_data ();  | 
|
| 74262 | 81  | 
                val _ = report {def = true} serial pos;
 | 
| 63543 | 82  | 
              in SOME {result = result, level = level, serial = serial, pos = pos} end
 | 
83  | 
          | SOME {level, serial, pos, ...} =>
 | 
|
| 74262 | 84  | 
              (report {def = false} serial pos;
 | 
| 63543 | 85  | 
                SOME {result = result, level = level, serial = serial, pos = pos})));
 | 
86  | 
in  | 
|
87  | 
state  | 
|
88  | 
|> (Proof.map_context o Context.proof_map o Data.map o apsnd) (K calculation)  | 
|
89  | 
|> Proof.map_context (Proof_Context.put_thms false (calculationN, calc))  | 
|
90  | 
end;  | 
|
| 6778 | 91  | 
|
92  | 
||
| 18637 | 93  | 
|
| 6778 | 94  | 
(** attributes **)  | 
95  | 
||
| 12379 | 96  | 
(* add/del rules *)  | 
97  | 
||
| 67627 | 98  | 
val trans_add =  | 
99  | 
Thm.declaration_attribute (Data.map o apfst o apfst o Item_Net.update o Thm.trim_context);  | 
|
100  | 
||
101  | 
val trans_del =  | 
|
102  | 
Thm.declaration_attribute (Data.map o apfst o apfst o Item_Net.remove);  | 
|
| 6778 | 103  | 
|
| 18637 | 104  | 
val sym_add =  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
105  | 
Thm.declaration_attribute (fn th =>  | 
| 74149 | 106  | 
(Data.map o apfst o apsnd) (Item_Net.update (Thm.trim_context th)) #>  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
107  | 
Thm.attribute_declaration (Context_Rules.elim_query NONE) th);  | 
| 33369 | 108  | 
|
| 18637 | 109  | 
val sym_del =  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
110  | 
Thm.declaration_attribute (fn th =>  | 
| 74149 | 111  | 
(Data.map o apfst o apsnd) (Item_Net.remove th) #>  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
112  | 
Thm.attribute_declaration Context_Rules.rule_del th);  | 
| 12379 | 113  | 
|
114  | 
||
| 18637 | 115  | 
(* symmetric *)  | 
| 12379 | 116  | 
|
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
117  | 
val symmetric =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
118  | 
Thm.rule_attribute [] (fn context => fn th =>  | 
| 74149 | 119  | 
let val ctxt = Context.proof_of context in  | 
120  | 
(case Seq.chop 2 (Drule.multi_resolves (SOME ctxt) [th] (get_sym_rules ctxt)) of  | 
|
121  | 
([th'], _) => Drule.zero_var_indexes th'  | 
|
122  | 
      | ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
 | 
|
123  | 
      | _ => raise THM ("symmetric: multiple unifiers", 1, [th]))
 | 
|
124  | 
end);  | 
|
| 12379 | 125  | 
|
| 6778 | 126  | 
|
127  | 
(* concrete syntax *)  | 
|
128  | 
||
| 53171 | 129  | 
val _ = Theory.setup  | 
| 67147 | 130  | 
(Attrib.setup \<^binding>\<open>trans\<close> (Attrib.add_del trans_add trans_del)  | 
| 30528 | 131  | 
"declaration of transitivity rule" #>  | 
| 67147 | 132  | 
Attrib.setup \<^binding>\<open>sym\<close> (Attrib.add_del sym_add sym_del)  | 
| 30528 | 133  | 
"declaration of symmetry rule" #>  | 
| 67147 | 134  | 
Attrib.setup \<^binding>\<open>symmetric\<close> (Scan.succeed symmetric)  | 
| 30528 | 135  | 
"resolution with symmetry rule" #>  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38332 
diff
changeset
 | 
136  | 
Global_Theory.add_thms  | 
| 29579 | 137  | 
[((Binding.empty, transitive_thm), [trans_add]),  | 
| 53171 | 138  | 
((Binding.empty, symmetric_thm), [sym_add])] #> snd);  | 
| 15801 | 139  | 
|
| 6778 | 140  | 
|
| 6787 | 141  | 
|
| 6778 | 142  | 
(** proof commands **)  | 
143  | 
||
| 
14555
 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 
wenzelm 
parents: 
14549 
diff
changeset
 | 
144  | 
fun assert_sane final =  | 
| 59126 | 145  | 
if final then Proof.assert_forward  | 
146  | 
else  | 
|
147  | 
Proof.assert_forward_or_chain #>  | 
|
148  | 
tap (fn state =>  | 
|
149  | 
if can Proof.assert_chain state then  | 
|
150  | 
Context_Position.report (Proof.context_of state) (Position.thread_data ()) Markup.improper  | 
|
151  | 
else ());  | 
|
| 
14555
 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 
wenzelm 
parents: 
14549 
diff
changeset
 | 
152  | 
|
| 
37047
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
153  | 
fun maintain_calculation int final calc state =  | 
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
154  | 
let  | 
| 63541 | 155  | 
val state' = state  | 
| 63543 | 156  | 
|> update_calculation (SOME calc)  | 
| 63541 | 157  | 
|> Proof.improper_reset_facts;  | 
| 
37047
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
158  | 
val ctxt' = Proof.context_of state';  | 
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
159  | 
val _ =  | 
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
160  | 
if int then  | 
| 56894 | 161  | 
Proof_Context.pretty_fact ctxt'  | 
162  | 
(Proof_Context.full_name ctxt' (Binding.name calculationN), calc)  | 
|
| 58843 | 163  | 
|> Pretty.string_of |> writeln  | 
| 
37047
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
164  | 
else ();  | 
| 63543 | 165  | 
in state' |> final ? (update_calculation NONE #> Proof.chain_facts calc) end;  | 
| 8562 | 166  | 
|
167  | 
||
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
168  | 
(* also and finally *)  | 
| 8562 | 169  | 
|
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
170  | 
fun calculate prep_rules final raw_rules int state =  | 
| 6778 | 171  | 
let  | 
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
172  | 
val ctxt = Proof.context_of state;  | 
| 61268 | 173  | 
val pretty_thm = Thm.pretty_thm ctxt;  | 
174  | 
val pretty_thm_item = Thm.pretty_thm_item ctxt;  | 
|
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
175  | 
|
| 12805 | 176  | 
val strip_assums_concl = Logic.strip_assums_concl o Thm.prop_of;  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
177  | 
val eq_prop = op aconv o apply2 (Envir.beta_eta_contract o strip_assums_concl);  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
178  | 
fun check_projection ths th =  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
179  | 
(case find_index (curry eq_prop th) ths of  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
180  | 
~1 => Seq.Result [th]  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
181  | 
| i =>  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
182  | 
Seq.Error (fn () =>  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
183  | 
(Pretty.string_of o Pretty.chunks)  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
184  | 
[Pretty.block [Pretty.str "Vacuous calculation result:", Pretty.brk 1, pretty_thm th],  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
185  | 
(Pretty.block o Pretty.fbreaks)  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
186  | 
                (Pretty.str ("derived as projection (" ^ string_of_int (i + 1) ^ ") from:") ::
 | 
| 51584 | 187  | 
map pretty_thm_item ths)]));  | 
| 8300 | 188  | 
|
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
189  | 
val opt_rules = Option.map (prep_rules ctxt) raw_rules;  | 
| 11097 | 190  | 
fun combine ths =  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
191  | 
Seq.append  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
192  | 
((case opt_rules of  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
193  | 
SOME rules => rules  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
194  | 
| NONE =>  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
195  | 
(case ths of  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
196  | 
[] => Item_Net.content (#1 (get_rules ctxt))  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
197  | 
| th :: _ => Item_Net.retrieve (#1 (get_rules ctxt)) (strip_assums_concl th)))  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
198  | 
|> Seq.of_list |> Seq.maps (Drule.multi_resolve (SOME ctxt) ths)  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
199  | 
|> Seq.map (check_projection ths))  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
200  | 
(Seq.single (Seq.Error (fn () =>  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
201  | 
(Pretty.string_of o Pretty.block o Pretty.fbreaks)  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
202  | 
(Pretty.str "No matching trans rules for calculation:" ::  | 
| 51584 | 203  | 
map pretty_thm_item ths))));  | 
| 7414 | 204  | 
|
| 
14555
 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 
wenzelm 
parents: 
14549 
diff
changeset
 | 
205  | 
val facts = Proof.the_facts (assert_sane final state);  | 
| 6903 | 206  | 
val (initial, calculations) =  | 
| 63543 | 207  | 
(case check state of  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
208  | 
NONE => (true, Seq.single (Seq.Result facts))  | 
| 
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
209  | 
| SOME calc => (false, combine (calc @ facts)));  | 
| 
37047
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
210  | 
|
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
211  | 
val _ = initial andalso final andalso error "No calculation yet";  | 
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
212  | 
val _ = initial andalso is_some opt_rules andalso  | 
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
213  | 
error "Initial calculation -- no rules to be given";  | 
| 6778 | 214  | 
in  | 
| 
49868
 
3039922ffd8d
more informative errors for 'also' and 'finally';
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
215  | 
calculations |> Seq.map_result (fn calc => maintain_calculation int final calc state)  | 
| 6778 | 216  | 
end;  | 
217  | 
||
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
218  | 
val also = calculate (K I) false;  | 
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
219  | 
val also_cmd = calculate Attrib.eval_thms false;  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
220  | 
val finally = calculate (K I) true;  | 
| 
38332
 
6551e310e7cb
prefer plain Attrib.eval_thmss -- also means the assert_forward of Proof.get_thmss_cmd is skipped, leading to uniform (albeit odd) behaviour concerning forward chaining;
 
wenzelm 
parents: 
37047 
diff
changeset
 | 
221  | 
val finally_cmd = calculate Attrib.eval_thms true;  | 
| 6778 | 222  | 
|
223  | 
||
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
224  | 
(* moreover and ultimately *)  | 
| 8562 | 225  | 
|
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
226  | 
fun collect final int state =  | 
| 8588 | 227  | 
let  | 
| 
14555
 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 
wenzelm 
parents: 
14549 
diff
changeset
 | 
228  | 
val facts = Proof.the_facts (assert_sane final state);  | 
| 8588 | 229  | 
val (initial, thms) =  | 
| 63543 | 230  | 
(case check state of  | 
| 15531 | 231  | 
NONE => (true, [])  | 
232  | 
| SOME thms => (false, thms));  | 
|
| 8588 | 233  | 
val calc = thms @ facts;  | 
| 
37047
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
234  | 
val _ = initial andalso final andalso error "No calculation yet";  | 
| 
 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 
wenzelm 
parents: 
36323 
diff
changeset
 | 
235  | 
in maintain_calculation int final calc state end;  | 
| 8588 | 236  | 
|
| 
17350
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
237  | 
val moreover = collect false;  | 
| 
 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 
wenzelm 
parents: 
16571 
diff
changeset
 | 
238  | 
val ultimately = collect true;  | 
| 8562 | 239  | 
|
| 6778 | 240  | 
end;  |