| author | wenzelm | 
| Wed, 29 Dec 2010 17:34:41 +0100 | |
| changeset 41413 | 64cd30d6b0b8 | 
| parent 39557 | fe5722fce758 | 
| child 42360 | da8817d01e7c | 
| 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: 
16571diff
changeset | 4 | Generic calculational proofs. | 
| 6778 | 5 | *) | 
| 6 | ||
| 7 | signature CALCULATION = | |
| 8 | sig | |
| 21506 | 9 | val print_rules: Proof.context -> unit | 
| 16571 | 10 | val get_calculation: 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 | |
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
33519diff
changeset | 16 | val also: thm list option -> bool -> Proof.state -> Proof.state Seq.seq | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 17 | val also_cmd: (Facts.ref * Attrib.src list) list option -> | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 18 | bool -> Proof.state -> Proof.state Seq.seq | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
33519diff
changeset | 19 | val finally: thm list option -> bool -> Proof.state -> Proof.state Seq.seq | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
33519diff
changeset | 20 | val finally_cmd: (Facts.ref * Attrib.src list) list option -> bool -> | 
| 22573 | 21 | Proof.state -> Proof.state Seq.seq | 
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 22 | val moreover: bool -> Proof.state -> Proof.state | 
| 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
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 | |
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 31 | structure Data = Generic_Data | 
| 18637 | 32 | ( | 
| 30560 | 33 | type T = (thm Item_Net.T * thm list) * (thm list * int) option; | 
| 34 | val empty = ((Thm.elim_rules, []), NONE); | |
| 18637 | 35 | val extend = I; | 
| 33519 | 36 | fun merge (((trans1, sym1), _), ((trans2, sym2), _)) = | 
| 30560 | 37 | ((Item_Net.merge (trans1, trans2), Thm.merge_thms (sym1, sym2)), NONE); | 
| 18637 | 38 | ); | 
| 39 | ||
| 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: 
37047diff
changeset | 40 | val get_rules = #1 o Data.get o Context.Proof; | 
| 
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: 
37047diff
changeset | 41 | |
| 22846 | 42 | fun print_rules 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: 
37047diff
changeset | 43 | let val (trans, sym) = get_rules ctxt in | 
| 22846 | 44 | [Pretty.big_list "transitivity rules:" | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30560diff
changeset | 45 | (map (Display.pretty_thm ctxt) (Item_Net.content trans)), | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30560diff
changeset | 46 | Pretty.big_list "symmetry rules:" (map (Display.pretty_thm ctxt) sym)] | 
| 22846 | 47 | |> Pretty.chunks |> Pretty.writeln | 
| 48 | end; | |
| 6778 | 49 | |
| 50 | ||
| 51 | (* access calculation *) | |
| 52 | ||
| 53 | fun get_calculation state = | |
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 54 | (case #2 (Data.get (Context.Proof (Proof.context_of state))) of | 
| 15531 | 55 | NONE => NONE | 
| 56 | | SOME (thms, lev) => if lev = Proof.level state then SOME thms else NONE); | |
| 6778 | 57 | |
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 58 | val calculationN = "calculation"; | 
| 6778 | 59 | |
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 60 | fun put_calculation calc = | 
| 18728 | 61 | `Proof.level #-> (fn lev => Proof.map_context (Context.proof_map | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 62 | (Data.map (apsnd (K (Option.map (rpair lev) calc)))))) | 
| 26252 | 63 | #> Proof.put_thms false (calculationN, calc); | 
| 6778 | 64 | |
| 65 | ||
| 18637 | 66 | |
| 6778 | 67 | (** attributes **) | 
| 68 | ||
| 12379 | 69 | (* add/del rules *) | 
| 70 | ||
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 71 | val trans_add = Thm.declaration_attribute (Data.map o apfst o apfst o Item_Net.update); | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 72 | val trans_del = Thm.declaration_attribute (Data.map o apfst o apfst o Item_Net.remove); | 
| 6778 | 73 | |
| 18637 | 74 | val sym_add = | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 75 | Thm.declaration_attribute (Data.map o apfst o apsnd o Thm.add_thm) | 
| 33369 | 76 | #> Context_Rules.elim_query NONE; | 
| 77 | ||
| 18637 | 78 | val sym_del = | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 79 | Thm.declaration_attribute (Data.map o apfst o apsnd o Thm.del_thm) | 
| 33369 | 80 | #> Context_Rules.rule_del; | 
| 12379 | 81 | |
| 82 | ||
| 18637 | 83 | (* symmetric *) | 
| 12379 | 84 | |
| 18728 | 85 | val symmetric = Thm.rule_attribute (fn x => fn th => | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 86 | (case Seq.chop 2 (Drule.multi_resolves [th] (#2 (#1 (Data.get x)))) of | 
| 20898 | 87 | ([th'], _) => Drule.zero_var_indexes th' | 
| 12379 | 88 |   | ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
 | 
| 89 |   | _ => raise THM ("symmetric: multiple unifiers", 1, [th])));
 | |
| 90 | ||
| 6778 | 91 | |
| 92 | (* concrete syntax *) | |
| 93 | ||
| 26463 | 94 | val _ = Context.>> (Context.map_theory | 
| 30528 | 95 | (Attrib.setup (Binding.name "trans") (Attrib.add_del trans_add trans_del) | 
| 96 | "declaration of transitivity rule" #> | |
| 97 | Attrib.setup (Binding.name "sym") (Attrib.add_del sym_add sym_del) | |
| 98 | "declaration of symmetry rule" #> | |
| 99 | Attrib.setup (Binding.name "symmetric") (Scan.succeed symmetric) | |
| 100 | "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: 
38332diff
changeset | 101 | Global_Theory.add_thms | 
| 29579 | 102 | [((Binding.empty, transitive_thm), [trans_add]), | 
| 103 | ((Binding.empty, symmetric_thm), [sym_add])] #> snd)); | |
| 15801 | 104 | |
| 6778 | 105 | |
| 6787 | 106 | |
| 6778 | 107 | (** proof commands **) | 
| 108 | ||
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 109 | fun assert_sane final = | 
| 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 110 | if final then Proof.assert_forward else Proof.assert_forward_or_chain; | 
| 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 111 | |
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 112 | fun maintain_calculation int final calc state = | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 113 | let | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 114 | val state' = put_calculation (SOME calc) state; | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 115 | val ctxt' = Proof.context_of state'; | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 116 | val _ = | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 117 | if int then | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 118 | Pretty.writeln | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 119 | (ProofContext.pretty_fact ctxt' | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 120 | (ProofContext.full_name ctxt' (Binding.name calculationN), calc)) | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 121 | else (); | 
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 122 | in state' |> final ? (put_calculation NONE #> Proof.chain_facts calc) end; | 
| 8562 | 123 | |
| 124 | ||
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 125 | (* also and finally *) | 
| 8562 | 126 | |
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 127 | fun calculate prep_rules final raw_rules int state = | 
| 6778 | 128 | 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: 
37047diff
changeset | 129 | val ctxt = Proof.context_of state; | 
| 
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: 
37047diff
changeset | 130 | |
| 12805 | 131 | val strip_assums_concl = Logic.strip_assums_concl o Thm.prop_of; | 
| 18947 | 132 | val eq_prop = op aconv o pairself (Envir.beta_eta_contract o strip_assums_concl); | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 133 | fun projection ths th = exists (curry eq_prop th) ths; | 
| 8300 | 134 | |
| 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: 
37047diff
changeset | 135 | val opt_rules = Option.map (prep_rules ctxt) raw_rules; | 
| 11097 | 136 | fun combine ths = | 
| 15531 | 137 | (case opt_rules of SOME rules => rules | 
| 138 | | NONE => | |
| 33373 | 139 | (case ths of | 
| 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: 
37047diff
changeset | 140 | [] => Item_Net.content (#1 (get_rules ctxt)) | 
| 
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: 
37047diff
changeset | 141 | | th :: _ => Item_Net.retrieve (#1 (get_rules ctxt)) (strip_assums_concl th))) | 
| 18223 | 142 | |> Seq.of_list |> Seq.maps (Drule.multi_resolve ths) | 
| 11097 | 143 | |> Seq.filter (not o projection ths); | 
| 7414 | 144 | |
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 145 | val facts = Proof.the_facts (assert_sane final state); | 
| 6903 | 146 | val (initial, calculations) = | 
| 6778 | 147 | (case get_calculation state of | 
| 15531 | 148 | NONE => (true, Seq.single facts) | 
| 149 | | SOME calc => (false, Seq.map single (combine (calc @ facts)))); | |
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 150 | |
| 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 151 | 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: 
36323diff
changeset | 152 | 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: 
36323diff
changeset | 153 | error "Initial calculation -- no rules to be given"; | 
| 6778 | 154 | in | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 155 | calculations |> Seq.map (fn calc => maintain_calculation int final calc state) | 
| 6778 | 156 | end; | 
| 157 | ||
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
33519diff
changeset | 158 | 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: 
37047diff
changeset | 159 | val also_cmd = calculate Attrib.eval_thms false; | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
33519diff
changeset | 160 | 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: 
37047diff
changeset | 161 | val finally_cmd = calculate Attrib.eval_thms true; | 
| 6778 | 162 | |
| 163 | ||
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 164 | (* moreover and ultimately *) | 
| 8562 | 165 | |
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 166 | fun collect final int state = | 
| 8588 | 167 | let | 
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 168 | val facts = Proof.the_facts (assert_sane final state); | 
| 8588 | 169 | val (initial, thms) = | 
| 170 | (case get_calculation state of | |
| 15531 | 171 | NONE => (true, []) | 
| 172 | | SOME thms => (false, thms)); | |
| 8588 | 173 | val calc = thms @ facts; | 
| 37047 
4a95a50d0ec4
print calculation result in the context where the fact is actually defined -- proper externing;
 wenzelm parents: 
36323diff
changeset | 174 | 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: 
36323diff
changeset | 175 | in maintain_calculation int final calc state end; | 
| 8588 | 176 | |
| 17350 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 177 | val moreover = collect false; | 
| 
26cd3756377a
more self-contained proof elements (material from isar_thy.ML);
 wenzelm parents: 
16571diff
changeset | 178 | val ultimately = collect true; | 
| 8562 | 179 | |
| 6778 | 180 | end; |