author | bulwahn |
Mon, 02 Jul 2012 12:23:30 +0200 | |
changeset 48179 | 18461f745b4a |
parent 47296 | c82a0b2606a1 |
child 49748 | a346daa8a1f4 |
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:
20049
diff
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:
26424
diff
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:
45406
diff
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:
45406
diff
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:
35739
diff
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:
42496
diff
changeset
|
57 |
val (bs, rhss) = split_list args; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
58 |
val Ts = map Term.fastype_of rhss; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
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:
42496
diff
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:
42501
diff
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:
42501
diff
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; |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
26424
diff
changeset
|
92 |
val ((bfacts, atts), rhss) = specs |> split_list |>> split_list; |
42496 | 93 |
val names = map2 Thm.def_binding_optional xs bfacts; |
20887 | 94 |
val eqs = mk_def ctxt (xs ~~ rhss); |
95 |
val lhss = map (fst o Logic.dest_equals) eqs; |
|
18830 | 96 |
in |
97 |
ctxt |
|
42496 | 98 |
|> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2 |
20980 | 99 |
|> fold Variable.declare_term eqs |
42360 | 100 |
|> Proof_Context.add_assms_i def_export |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
26424
diff
changeset
|
101 |
(map2 (fn a => fn eq => (a, [(eq, [])])) (names ~~ atts) eqs) |
20887 | 102 |
|>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss |
18830 | 103 |
end; |
104 |
||
25104 | 105 |
fun add_def (var, rhs) ctxt = |
30211 | 106 |
let val ([(lhs, (_, th))], ctxt') = add_defs [(var, (Thm.empty_binding, rhs))] ctxt |
25104 | 107 |
in ((lhs, th), ctxt') end; |
25025 | 108 |
|
18840 | 109 |
|
25119 | 110 |
(* fixed_abbrev *) |
111 |
||
112 |
fun fixed_abbrev ((x, mx), rhs) ctxt = |
|
113 |
let |
|
114 |
val T = Term.fastype_of rhs; |
|
115 |
val ([x'], ctxt') = ctxt |
|
116 |
|> Variable.declare_term rhs |
|
42360 | 117 |
|> Proof_Context.add_fixes [(x, SOME T, mx)]; |
25119 | 118 |
val lhs = Free (x', T); |
119 |
val _ = cert_def ctxt' (Logic.mk_equals (lhs, rhs)); |
|
120 |
fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]); |
|
121 |
val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt'; |
|
122 |
in ((lhs, rhs), ctxt'') end; |
|
123 |
||
124 |
||
24985 | 125 |
(* specific export -- result based on educated guessing *) |
21568 | 126 |
|
24985 | 127 |
(* |
128 |
[xs, xs == as] |
|
129 |
: |
|
130 |
B xs |
|
131 |
-------------- |
|
132 |
B as |
|
133 |
*) |
|
134 |
fun export inner outer = (*beware of closure sizes*) |
|
21568 | 135 |
let |
22671 | 136 |
val exp = Assumption.export false inner outer; |
45406
b24ecaa46edb
clarified Local_Defs.export: avoid costly still_fixed test, return all defs;
wenzelm
parents:
42501
diff
changeset
|
137 |
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:
47238
diff
changeset
|
138 |
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:
45406
diff
changeset
|
139 |
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:
45406
diff
changeset
|
140 |
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:
45406
diff
changeset
|
141 |
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:
45406
diff
changeset
|
142 |
val th' = exp th; |
47294
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
wenzelm
parents:
47238
diff
changeset
|
143 |
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:
47238
diff
changeset
|
144 |
(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:
47238
diff
changeset
|
145 |
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:
45406
diff
changeset
|
146 |
| 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:
45406
diff
changeset
|
147 |
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:
45406
diff
changeset
|
148 |
(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:
47238
diff
changeset
|
149 |
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:
45406
diff
changeset
|
150 |
| SOME u => |
47294
008b7858f3c0
retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
wenzelm
parents:
47238
diff
changeset
|
151 |
if t aconv u then (asm, false) |
47296 | 152 |
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:
47238
diff
changeset
|
153 |
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:
45406
diff
changeset
|
154 |
in (pairself (map #1) (List.partition #2 defs_asms), th') end |
22671 | 155 |
end; |
21568 | 156 |
|
24985 | 157 |
(* |
158 |
[xs, xs == as] |
|
159 |
: |
|
160 |
TERM b xs |
|
161 |
-------------- and -------------- |
|
162 |
TERM b as b xs == b as |
|
163 |
*) |
|
164 |
fun export_cterm inner outer ct = |
|
35717 | 165 |
export inner outer (Drule.mk_term ct) ||> Drule.dest_term; |
24985 | 166 |
|
47238 | 167 |
fun contract defs ct th = |
168 |
th COMP (Raw_Simplifier.rewrite true defs ct COMP_INCR Drule.equal_elim_rule2); |
|
21844 | 169 |
|
18840 | 170 |
|
47237
b61a11dea74c
more precise Local_Defs.expand wrt. *local* prems only;
wenzelm
parents:
47060
diff
changeset
|
171 |
|
18840 | 172 |
(** defived definitions **) |
173 |
||
174 |
(* transformation rules *) |
|
175 |
||
33519 | 176 |
structure Rules = Generic_Data |
18840 | 177 |
( |
178 |
type T = thm list; |
|
33519 | 179 |
val empty = []; |
18840 | 180 |
val extend = I; |
33519 | 181 |
val merge = Thm.merge_thms; |
18840 | 182 |
); |
183 |
||
22846 | 184 |
fun print_rules ctxt = |
185 |
Pretty.writeln (Pretty.big_list "definitional transformations:" |
|
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:
31794
diff
changeset
|
186 |
(map (Display.pretty_thm ctxt) (Rules.get (Context.Proof ctxt)))); |
18840 | 187 |
|
24039
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23541
diff
changeset
|
188 |
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:
23541
diff
changeset
|
189 |
val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm); |
18840 | 190 |
|
191 |
||
18875 | 192 |
(* meta rewrite rules *) |
18840 | 193 |
|
23541 | 194 |
fun meta_rewrite_conv ctxt = |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
195 |
Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
196 |
(Raw_Simplifier.context ctxt empty_ss |
45620
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45407
diff
changeset
|
197 |
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:
45407
diff
changeset
|
198 |
|> Raw_Simplifier.add_eqcong Drule.equals_cong); (*protect meta-level equality*) |
18840 | 199 |
|
23541 | 200 |
val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv; |
18840 | 201 |
|
18875 | 202 |
|
203 |
(* rewriting with object-level rules *) |
|
204 |
||
21506 | 205 |
fun meta f ctxt = f o map (meta_rewrite_rule ctxt); |
18840 | 206 |
|
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
207 |
val unfold = meta Raw_Simplifier.rewrite_rule; |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
208 |
val unfold_goals = meta Raw_Simplifier.rewrite_goals_rule; |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
209 |
val unfold_tac = meta Raw_Simplifier.rewrite_goals_tac; |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
210 |
val fold = meta Raw_Simplifier.fold_rule; |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
211 |
val fold_tac = meta Raw_Simplifier.fold_goals_tac; |
18840 | 212 |
|
213 |
||
214 |
(* derived defs -- potentially within the object-logic *) |
|
215 |
||
18950 | 216 |
fun derived_def ctxt conditional prop = |
18840 | 217 |
let |
218 |
val ((c, T), rhs) = prop |
|
42360 | 219 |
|> Thm.cterm_of (Proof_Context.theory_of ctxt) |
23541 | 220 |
|> meta_rewrite_conv ctxt |
18840 | 221 |
|> (snd o Logic.dest_equals o Thm.prop_of) |
21568 | 222 |
|> conditional ? Logic.strip_imp_concl |
18950 | 223 |
|> (abs_def o #2 o cert_def ctxt); |
20909 | 224 |
fun prove ctxt' def = |
42482 | 225 |
Goal.prove ctxt' (Variable.add_free_names ctxt' prop []) [] prop (K (ALLGOALS |
226 |
(CONVERSION (meta_rewrite_conv ctxt') THEN' |
|
227 |
rewrite_goal_tac [def] THEN' |
|
228 |
resolve_tac [Drule.reflexive_thm]))) |
|
229 |
handle ERROR msg => cat_error msg "Failed to prove definitional specification"; |
|
18840 | 230 |
in (((c, T), rhs), prove) end; |
231 |
||
18830 | 232 |
end; |