| author | blanchet | 
| Tue, 23 Apr 2013 17:13:14 +0200 | |
| changeset 51744 | 0468af6546ff | 
| parent 51717 | 9e7d1c139569 | 
| child 54742 | 7a86358a3c0b | 
| permissions | -rw-r--r-- | 
| 18830 | 1 | (* Title: Pure/Isar/local_defs.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 18840 | 4 | Local definitions. | 
| 18830 | 5 | *) | 
| 6 | ||
| 7 | signature LOCAL_DEFS = | |
| 8 | sig | |
| 20306 | 9 | val cert_def: Proof.context -> term -> (string * typ) * term | 
| 18830 | 10 | val abs_def: term -> (string * typ) * term | 
| 21684 | 11 | val expand: cterm list -> thm -> thm | 
| 20224 
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
 wenzelm parents: 
20049diff
changeset | 12 | val def_export: Assumption.export | 
| 30211 | 13 | val add_defs: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context -> | 
| 14 | (term * (string * thm)) list * Proof.context | |
| 29581 | 15 | val add_def: (binding * mixfix) * term -> Proof.context -> (term * thm) * Proof.context | 
| 16 | val fixed_abbrev: (binding * mixfix) * term -> Proof.context -> | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
26424diff
changeset | 17 | (term * term) * Proof.context | 
| 45407 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 18 | val export: Proof.context -> Proof.context -> thm -> (thm list * thm list) * thm | 
| 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 19 | val export_cterm: Proof.context -> Proof.context -> cterm -> (thm list * thm list) * cterm | 
| 47238 | 20 | val contract: thm list -> cterm -> thm -> thm | 
| 21506 | 21 | val print_rules: Proof.context -> unit | 
| 18840 | 22 | val defn_add: attribute | 
| 23 | val defn_del: attribute | |
| 23541 | 24 | val meta_rewrite_conv: Proof.context -> conv | 
| 21506 | 25 | val meta_rewrite_rule: Proof.context -> thm -> thm | 
| 20306 | 26 | val unfold: Proof.context -> thm list -> thm -> thm | 
| 27 | val unfold_goals: Proof.context -> thm list -> thm -> thm | |
| 28 | val unfold_tac: Proof.context -> thm list -> tactic | |
| 29 | val fold: Proof.context -> thm list -> thm -> thm | |
| 30 | val fold_tac: Proof.context -> thm list -> tactic | |
| 31 | val derived_def: Proof.context -> bool -> term -> | |
| 20909 | 32 | ((string * typ) * term) * (Proof.context -> thm -> thm) | 
| 18830 | 33 | end; | 
| 34 | ||
| 35624 | 35 | structure Local_Defs: LOCAL_DEFS = | 
| 18830 | 36 | struct | 
| 37 | ||
| 18840 | 38 | (** primitive definitions **) | 
| 39 | ||
| 18830 | 40 | (* prepare defs *) | 
| 41 | ||
| 42 | fun cert_def ctxt eq = | |
| 43 | let | |
| 40242 | 44 | fun err msg = | 
| 45 |       cat_error msg ("The error(s) above occurred in definition:\n" ^
 | |
| 46 | quote (Syntax.string_of_term ctxt eq)); | |
| 18950 | 47 | val ((lhs, _), eq') = eq | 
| 39133 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 wenzelm parents: 
35739diff
changeset | 48 | |> Sign.no_vars ctxt | 
| 33385 | 49 | |> Primitive_Defs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true) | 
| 18950 | 50 | handle TERM (msg, _) => err msg | ERROR msg => err msg; | 
| 51 | in (Term.dest_Free (Term.head_of lhs), eq') end; | |
| 18830 | 52 | |
| 33385 | 53 | val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free; | 
| 18830 | 54 | |
| 18875 | 55 | fun mk_def ctxt args = | 
| 56 | let | |
| 42501 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 wenzelm parents: 
42496diff
changeset | 57 | val (bs, rhss) = split_list args; | 
| 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 wenzelm parents: 
42496diff
changeset | 58 | val Ts = map Term.fastype_of rhss; | 
| 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 wenzelm parents: 
42496diff
changeset | 59 | val (xs, _) = Proof_Context.add_fixes (map2 (fn b => fn T => (b, SOME T, NoSyn)) bs Ts) ctxt; | 
| 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 wenzelm parents: 
42496diff
changeset | 60 | val lhss = ListPair.map Free (xs, Ts); | 
| 18875 | 61 | in map Logic.mk_equals (lhss ~~ rhss) end; | 
| 62 | ||
| 18830 | 63 | |
| 64 | (* export defs *) | |
| 65 | ||
| 20021 | 66 | val head_of_def = | 
| 45406 
b24ecaa46edb
clarified Local_Defs.export: avoid costly still_fixed test, return all defs;
 wenzelm parents: 
42501diff
changeset | 67 | Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body; | 
| 20021 | 68 | |
| 18830 | 69 | |
| 18875 | 70 | (* | 
| 20887 | 71 | [x, x == a] | 
| 18896 | 72 | : | 
| 73 | B x | |
| 74 | ----------- | |
| 20887 | 75 | B a | 
| 18875 | 76 | *) | 
| 21684 | 77 | fun expand defs = | 
| 78 | Drule.implies_intr_list defs | |
| 45406 
b24ecaa46edb
clarified Local_Defs.export: avoid costly still_fixed test, return all defs;
 wenzelm parents: 
42501diff
changeset | 79 | #> Drule.generalize ([], map (#1 o head_of_def o Thm.term_of) defs) | 
| 21684 | 80 | #> funpow (length defs) (fn th => Drule.reflexive_thm RS th); | 
| 81 | ||
| 21801 | 82 | val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of); | 
| 21684 | 83 | |
| 84 | fun def_export _ defs = (expand defs, expand_term defs); | |
| 18830 | 85 | |
| 86 | ||
| 87 | (* add defs *) | |
| 88 | ||
| 20887 | 89 | fun add_defs defs ctxt = | 
| 18830 | 90 | let | 
| 42496 | 91 | val ((xs, mxs), specs) = defs |> split_list |>> split_list; | 
| 49748 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
47296diff
changeset | 92 | val (bs, rhss) = specs |> split_list; | 
| 20887 | 93 | val eqs = mk_def ctxt (xs ~~ rhss); | 
| 94 | val lhss = map (fst o Logic.dest_equals) eqs; | |
| 18830 | 95 | in | 
| 96 | ctxt | |
| 42496 | 97 | |> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2 | 
| 20980 | 98 | |> fold Variable.declare_term eqs | 
| 49748 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
47296diff
changeset | 99 | |> Proof_Context.add_assms_i def_export (map2 (fn b => fn eq => (b, [(eq, [])])) bs eqs) | 
| 20887 | 100 | |>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss | 
| 18830 | 101 | end; | 
| 102 | ||
| 25104 | 103 | fun add_def (var, rhs) ctxt = | 
| 30211 | 104 | let val ([(lhs, (_, th))], ctxt') = add_defs [(var, (Thm.empty_binding, rhs))] ctxt | 
| 25104 | 105 | in ((lhs, th), ctxt') end; | 
| 25025 | 106 | |
| 18840 | 107 | |
| 25119 | 108 | (* fixed_abbrev *) | 
| 109 | ||
| 110 | fun fixed_abbrev ((x, mx), rhs) ctxt = | |
| 111 | let | |
| 112 | val T = Term.fastype_of rhs; | |
| 113 | val ([x'], ctxt') = ctxt | |
| 114 | |> Variable.declare_term rhs | |
| 42360 | 115 | |> Proof_Context.add_fixes [(x, SOME T, mx)]; | 
| 25119 | 116 | val lhs = Free (x', T); | 
| 117 | val _ = cert_def ctxt' (Logic.mk_equals (lhs, rhs)); | |
| 118 | fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]); | |
| 119 | val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt'; | |
| 120 | in ((lhs, rhs), ctxt'') end; | |
| 121 | ||
| 122 | ||
| 24985 | 123 | (* specific export -- result based on educated guessing *) | 
| 21568 | 124 | |
| 24985 | 125 | (* | 
| 126 | [xs, xs == as] | |
| 127 | : | |
| 128 | B xs | |
| 129 | -------------- | |
| 130 | B as | |
| 131 | *) | |
| 132 | fun export inner outer = (*beware of closure sizes*) | |
| 21568 | 133 | let | 
| 22671 | 134 | val exp = Assumption.export false inner outer; | 
| 45406 
b24ecaa46edb
clarified Local_Defs.export: avoid costly still_fixed test, return all defs;
 wenzelm parents: 
42501diff
changeset | 135 | val exp_term = Assumption.export_term inner outer; | 
| 47294 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 136 | val asms = Assumption.local_assms_of inner outer; | 
| 45407 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 137 | in | 
| 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 138 | fn th => | 
| 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 139 | let | 
| 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 140 | val th' = exp th; | 
| 47294 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 141 | val defs_asms = asms |> map (Thm.assume #> (fn asm => | 
| 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 142 | (case try (head_of_def o Thm.prop_of) asm of | 
| 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 143 | NONE => (asm, false) | 
| 45407 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 144 | | SOME x => | 
| 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 145 | let val t = Free x in | 
| 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 146 | (case try exp_term t of | 
| 47294 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 147 | NONE => (asm, false) | 
| 45407 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 148 | | SOME u => | 
| 47294 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 149 | if t aconv u then (asm, false) | 
| 47296 | 150 | else (Drule.abs_def (Drule.gen_all asm), true)) | 
| 47294 
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
 wenzelm parents: 
47238diff
changeset | 151 | end))); | 
| 45407 
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
 wenzelm parents: 
45406diff
changeset | 152 | in (pairself (map #1) (List.partition #2 defs_asms), th') end | 
| 22671 | 153 | end; | 
| 21568 | 154 | |
| 24985 | 155 | (* | 
| 156 | [xs, xs == as] | |
| 157 | : | |
| 158 | TERM b xs | |
| 159 | -------------- and -------------- | |
| 160 | TERM b as b xs == b as | |
| 161 | *) | |
| 162 | fun export_cterm inner outer ct = | |
| 35717 | 163 | export inner outer (Drule.mk_term ct) ||> Drule.dest_term; | 
| 24985 | 164 | |
| 47238 | 165 | fun contract defs ct th = | 
| 166 | th COMP (Raw_Simplifier.rewrite true defs ct COMP_INCR Drule.equal_elim_rule2); | |
| 21844 | 167 | |
| 18840 | 168 | |
| 47237 
b61a11dea74c
more precise Local_Defs.expand wrt. *local* prems only;
 wenzelm parents: 
47060diff
changeset | 169 | |
| 18840 | 170 | (** defived definitions **) | 
| 171 | ||
| 51585 | 172 | (* transformation via rewrite rules *) | 
| 18840 | 173 | |
| 33519 | 174 | structure Rules = Generic_Data | 
| 18840 | 175 | ( | 
| 176 | type T = thm list; | |
| 33519 | 177 | val empty = []; | 
| 18840 | 178 | val extend = I; | 
| 33519 | 179 | val merge = Thm.merge_thms; | 
| 18840 | 180 | ); | 
| 181 | ||
| 22846 | 182 | fun print_rules ctxt = | 
| 51585 | 183 | Pretty.writeln (Pretty.big_list "definitional rewrite rules:" | 
| 51584 | 184 | (map (Display.pretty_thm_item ctxt) (Rules.get (Context.Proof ctxt)))); | 
| 18840 | 185 | |
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23541diff
changeset | 186 | val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23541diff
changeset | 187 | val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm); | 
| 18840 | 188 | |
| 189 | ||
| 18875 | 190 | (* meta rewrite rules *) | 
| 18840 | 191 | |
| 23541 | 192 | fun meta_rewrite_conv ctxt = | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40242diff
changeset | 193 | Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51585diff
changeset | 194 | (empty_simpset ctxt | 
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45407diff
changeset | 195 | addsimps (Rules.get (Context.Proof ctxt)) | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45407diff
changeset | 196 | |> Raw_Simplifier.add_eqcong Drule.equals_cong); (*protect meta-level equality*) | 
| 18840 | 197 | |
| 23541 | 198 | val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv; | 
| 18840 | 199 | |
| 18875 | 200 | |
| 201 | (* rewriting with object-level rules *) | |
| 202 | ||
| 21506 | 203 | fun meta f ctxt = f o map (meta_rewrite_rule ctxt); | 
| 18840 | 204 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40242diff
changeset | 205 | val unfold = meta Raw_Simplifier.rewrite_rule; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40242diff
changeset | 206 | val unfold_goals = meta Raw_Simplifier.rewrite_goals_rule; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40242diff
changeset | 207 | val unfold_tac = meta Raw_Simplifier.rewrite_goals_tac; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40242diff
changeset | 208 | val fold = meta Raw_Simplifier.fold_rule; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40242diff
changeset | 209 | val fold_tac = meta Raw_Simplifier.fold_goals_tac; | 
| 18840 | 210 | |
| 211 | ||
| 212 | (* derived defs -- potentially within the object-logic *) | |
| 213 | ||
| 18950 | 214 | fun derived_def ctxt conditional prop = | 
| 18840 | 215 | let | 
| 216 | val ((c, T), rhs) = prop | |
| 42360 | 217 | |> Thm.cterm_of (Proof_Context.theory_of ctxt) | 
| 23541 | 218 | |> meta_rewrite_conv ctxt | 
| 18840 | 219 | |> (snd o Logic.dest_equals o Thm.prop_of) | 
| 21568 | 220 | |> conditional ? Logic.strip_imp_concl | 
| 18950 | 221 | |> (abs_def o #2 o cert_def ctxt); | 
| 20909 | 222 | fun prove ctxt' def = | 
| 42482 | 223 | Goal.prove ctxt' (Variable.add_free_names ctxt' prop []) [] prop (K (ALLGOALS | 
| 224 | (CONVERSION (meta_rewrite_conv ctxt') THEN' | |
| 225 | rewrite_goal_tac [def] THEN' | |
| 226 | resolve_tac [Drule.reflexive_thm]))) | |
| 227 | handle ERROR msg => cat_error msg "Failed to prove definitional specification"; | |
| 18840 | 228 | in (((c, T), rhs), prove) end; | 
| 229 | ||
| 18830 | 230 | end; |