| author | wenzelm | 
| Wed, 29 Jun 2005 15:13:38 +0200 | |
| changeset 16608 | 4f8d7b83c7e2 | 
| parent 16571 | c1f41c98fd3c | 
| child 17350 | 26cd3756377a | 
| permissions | -rw-r--r-- | 
| 6778 | 1 | (* Title: Pure/Isar/calculation.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | Support for calculational proofs. | |
| 6 | *) | |
| 7 | ||
| 8 | signature CALCULATION = | |
| 9 | sig | |
| 16571 | 10 | val get_calculation: Proof.state -> thm list option | 
| 6778 | 11 | val print_global_rules: theory -> unit | 
| 12 | val print_local_rules: Proof.context -> unit | |
| 13 | val trans_add_global: theory attribute | |
| 14 | val trans_del_global: theory attribute | |
| 15 | val trans_add_local: Proof.context attribute | |
| 16 | val trans_del_local: Proof.context attribute | |
| 12379 | 17 | val sym_add_global: theory attribute | 
| 18 | val sym_del_global: theory attribute | |
| 19 | val sym_add_local: Proof.context attribute | |
| 20 | val sym_del_local: Proof.context attribute | |
| 21 | val symmetric_global: theory attribute | |
| 22 | val symmetric_local: Proof.context attribute | |
| 12055 | 23 | val also: thm list option -> (Proof.context -> thm list -> unit) | 
| 24 | -> Proof.state -> Proof.state Seq.seq | |
| 25 | val finally: thm list option -> (Proof.context -> thm list -> unit) | |
| 26 | -> Proof.state -> Proof.state Seq.seq | |
| 27 | val moreover: (Proof.context -> thm list -> unit) -> Proof.state -> Proof.state | |
| 28 | val ultimately: (Proof.context -> thm list -> unit) -> Proof.state -> Proof.state | |
| 6778 | 29 | end; | 
| 30 | ||
| 31 | structure Calculation: CALCULATION = | |
| 32 | struct | |
| 33 | ||
| 34 | (** global and local calculation data **) | |
| 35 | ||
| 16424 | 36 | (* global calculation *) | 
| 6778 | 37 | |
| 12379 | 38 | fun print_rules prt x (trans, sym) = | 
| 39 | [Pretty.big_list "transitivity rules:" (map (prt x) (NetRules.rules trans)), | |
| 40 | Pretty.big_list "symmetry rules:" (map (prt x) sym)] | |
| 41 | |> Pretty.chunks |> Pretty.writeln; | |
| 6778 | 42 | |
| 16424 | 43 | structure GlobalCalculation = TheoryDataFun | 
| 44 | (struct | |
| 6778 | 45 | val name = "Isar/calculation"; | 
| 12379 | 46 | type T = thm NetRules.T * thm list | 
| 6778 | 47 | |
| 12379 | 48 | val empty = (NetRules.elim, []); | 
| 6778 | 49 | val copy = I; | 
| 16424 | 50 | val extend = I; | 
| 51 | fun merge _ ((trans1, sym1), (trans2, sym2)) = | |
| 12379 | 52 | (NetRules.merge (trans1, trans2), Drule.merge_rules (sym1, sym2)); | 
| 12055 | 53 | val print = print_rules Display.pretty_thm_sg; | 
| 16424 | 54 | end); | 
| 6778 | 55 | |
| 15801 | 56 | val _ = Context.add_setup [GlobalCalculation.init]; | 
| 6778 | 57 | val print_global_rules = GlobalCalculation.print; | 
| 58 | ||
| 59 | ||
| 16424 | 60 | (* local calculation *) | 
| 6778 | 61 | |
| 16424 | 62 | structure LocalCalculation = ProofDataFun | 
| 63 | (struct | |
| 6778 | 64 | val name = "Isar/calculation"; | 
| 12379 | 65 | type T = (thm NetRules.T * thm list) * (thm list * int) option; | 
| 6778 | 66 | |
| 15531 | 67 | fun init thy = (GlobalCalculation.get thy, NONE); | 
| 12055 | 68 | fun print ctxt (rs, _) = print_rules ProofContext.pretty_thm ctxt rs; | 
| 16424 | 69 | end); | 
| 6778 | 70 | |
| 15801 | 71 | val _ = Context.add_setup [LocalCalculation.init]; | 
| 13371 | 72 | val get_local_rules = #1 o LocalCalculation.get o Proof.context_of; | 
| 6778 | 73 | val print_local_rules = LocalCalculation.print; | 
| 74 | ||
| 75 | ||
| 76 | (* access calculation *) | |
| 77 | ||
| 78 | fun get_calculation state = | |
| 13371 | 79 | (case #2 (LocalCalculation.get (Proof.context_of state)) of | 
| 15531 | 80 | NONE => NONE | 
| 81 | | SOME (thms, lev) => if lev = Proof.level state then SOME thms else NONE); | |
| 6778 | 82 | |
| 7414 | 83 | fun put_calculation thms state = | 
| 13371 | 84 | Proof.map_context | 
| 15531 | 85 | (LocalCalculation.put (get_local_rules state, SOME (thms, Proof.level state))) state; | 
| 6778 | 86 | |
| 6787 | 87 | fun reset_calculation state = | 
| 15531 | 88 | Proof.map_context (LocalCalculation.put (get_local_rules state, NONE)) state; | 
| 6787 | 89 | |
| 6778 | 90 | |
| 91 | ||
| 92 | (** attributes **) | |
| 93 | ||
| 12379 | 94 | (* add/del rules *) | 
| 95 | ||
| 96 | fun global_att f (x, thm) = (GlobalCalculation.map (f thm) x, thm); | |
| 97 | fun local_att f (x, thm) = (LocalCalculation.map (apfst (f thm)) x, thm); | |
| 6778 | 98 | |
| 12379 | 99 | val trans_add_global = global_att (apfst o NetRules.insert); | 
| 100 | val trans_del_global = global_att (apfst o NetRules.delete); | |
| 101 | val trans_add_local = local_att (apfst o NetRules.insert); | |
| 102 | val trans_del_local = local_att (apfst o NetRules.delete); | |
| 6778 | 103 | |
| 15531 | 104 | val sym_add_global = global_att (apsnd o Drule.add_rule) o ContextRules.elim_query_global NONE; | 
| 12402 | 105 | val sym_del_global = global_att (apsnd o Drule.del_rule) o ContextRules.rule_del_global; | 
| 15531 | 106 | val sym_add_local = local_att (apsnd o Drule.add_rule) o ContextRules.elim_query_local NONE; | 
| 12402 | 107 | val sym_del_local = local_att (apsnd o Drule.del_rule) o ContextRules.rule_del_local; | 
| 12379 | 108 | |
| 109 | ||
| 110 | (* symmetry *) | |
| 111 | ||
| 112 | fun gen_symmetric get_sym = Drule.rule_attribute (fn x => fn th => | |
| 113 | (case Seq.chop (2, Method.multi_resolves [th] (get_sym x)) of | |
| 114 | ([th'], _) => th' | |
| 115 |   | ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
 | |
| 116 |   | _ => raise THM ("symmetric: multiple unifiers", 1, [th])));
 | |
| 117 | ||
| 118 | val symmetric_global = gen_symmetric (#2 o GlobalCalculation.get); | |
| 119 | val symmetric_local = gen_symmetric (#2 o #1 o LocalCalculation.get); | |
| 6778 | 120 | |
| 121 | ||
| 122 | (* concrete syntax *) | |
| 123 | ||
| 124 | val trans_attr = | |
| 8634 | 125 | (Attrib.add_del_args trans_add_global trans_del_global, | 
| 126 | Attrib.add_del_args trans_add_local trans_del_local); | |
| 6778 | 127 | |
| 12379 | 128 | val sym_attr = | 
| 129 | (Attrib.add_del_args sym_add_global sym_del_global, | |
| 130 | Attrib.add_del_args sym_add_local sym_del_local); | |
| 131 | ||
| 15801 | 132 | val _ = Context.add_setup | 
| 133 | [Attrib.add_attributes | |
| 134 |    [("trans", trans_attr, "declaration of transitivity rule"),
 | |
| 135 |     ("sym", sym_attr, "declaration of symmetry rule"),
 | |
| 136 |     ("symmetric", (Attrib.no_args symmetric_global, Attrib.no_args symmetric_local),
 | |
| 137 | "resolution with symmetry rule")], | |
| 138 | #1 o PureThy.add_thms | |
| 139 |    [(("", transitive_thm), [trans_add_global]),
 | |
| 140 |     (("", symmetric_thm), [sym_add_global])]];
 | |
| 141 | ||
| 6778 | 142 | |
| 6787 | 143 | |
| 6778 | 144 | (** proof commands **) | 
| 145 | ||
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 146 | fun assert_sane final = | 
| 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 147 | 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 | 148 | |
| 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 149 | |
| 8588 | 150 | (* maintain calculation register *) | 
| 8562 | 151 | |
| 6778 | 152 | val calculationN = "calculation"; | 
| 153 | ||
| 8588 | 154 | fun maintain_calculation false calc state = | 
| 155 | state | |
| 156 | |> put_calculation calc | |
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 157 | |> Proof.put_thms (calculationN, calc) | 
| 8588 | 158 | | maintain_calculation true calc state = | 
| 159 | state | |
| 160 | |> reset_calculation | |
| 161 | |> Proof.reset_thms calculationN | |
| 14564 | 162 | |> Proof.simple_note_thms "" calc | 
| 8588 | 163 | |> Proof.chain; | 
| 8562 | 164 | |
| 165 | ||
| 166 | (* 'also' and 'finally' *) | |
| 167 | ||
| 8588 | 168 | fun err_if state b msg = if b then raise Proof.STATE (msg, state) else (); | 
| 169 | ||
| 6877 | 170 | fun calculate final opt_rules print state = | 
| 6778 | 171 | let | 
| 12805 | 172 | val strip_assums_concl = Logic.strip_assums_concl o Thm.prop_of; | 
| 9322 | 173 | val eq_prop = op aconv o pairself (Pattern.eta_contract o strip_assums_concl); | 
| 11097 | 174 | fun projection ths th = Library.exists (Library.curry eq_prop th) ths; | 
| 8300 | 175 | |
| 11097 | 176 | fun combine ths = | 
| 15531 | 177 | (case opt_rules of SOME rules => rules | 
| 178 | | NONE => | |
| 12379 | 179 | (case ths of [] => NetRules.rules (#1 (get_local_rules state)) | 
| 180 | | th :: _ => NetRules.retrieve (#1 (get_local_rules state)) (strip_assums_concl th))) | |
| 11097 | 181 | |> Seq.of_list |> Seq.map (Method.multi_resolve ths) |> Seq.flat | 
| 182 | |> Seq.filter (not o projection ths); | |
| 7414 | 183 | |
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 184 | val facts = Proof.the_facts (assert_sane final state); | 
| 6903 | 185 | val (initial, calculations) = | 
| 6778 | 186 | (case get_calculation state of | 
| 15531 | 187 | NONE => (true, Seq.single facts) | 
| 188 | | SOME calc => (false, Seq.map single (combine (calc @ facts)))); | |
| 6778 | 189 | in | 
| 8588 | 190 | err_if state (initial andalso final) "No calculation yet"; | 
| 15973 | 191 | err_if state (initial andalso is_some opt_rules) "Initial calculation -- no rules to be given"; | 
| 12055 | 192 | calculations |> Seq.map (fn calc => (print (Proof.context_of state) calc; | 
| 193 | state |> maintain_calculation final calc)) | |
| 6778 | 194 | end; | 
| 195 | ||
| 6782 | 196 | fun also print = calculate false print; | 
| 197 | fun finally print = calculate true print; | |
| 6778 | 198 | |
| 199 | ||
| 8588 | 200 | (* 'moreover' and 'ultimately' *) | 
| 8562 | 201 | |
| 8588 | 202 | fun collect final print state = | 
| 203 | let | |
| 14555 
341908d6c792
'also'/'moreover': do not interfere with current facts, allow in chain mode;
 wenzelm parents: 
14549diff
changeset | 204 | val facts = Proof.the_facts (assert_sane final state); | 
| 8588 | 205 | val (initial, thms) = | 
| 206 | (case get_calculation state of | |
| 15531 | 207 | NONE => (true, []) | 
| 208 | | SOME thms => (false, thms)); | |
| 8588 | 209 | val calc = thms @ facts; | 
| 210 | in | |
| 211 | err_if state (initial andalso final) "No calculation yet"; | |
| 12055 | 212 | print (Proof.context_of state) calc; | 
| 8588 | 213 | state |> maintain_calculation final calc | 
| 214 | end; | |
| 215 | ||
| 216 | fun moreover print = collect false print; | |
| 217 | fun ultimately print = collect true print; | |
| 8562 | 218 | |
| 219 | ||
| 6778 | 220 | end; |