author | wenzelm |
Mon, 07 Nov 2011 17:00:23 +0100 | |
changeset 45390 | e29521ef9059 |
parent 42501 | 2b8c34f53388 |
child 45406 | b24ecaa46edb |
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 |
21591 | 18 |
val export: Proof.context -> Proof.context -> thm -> thm list * thm |
35717 | 19 |
val export_cterm: Proof.context -> Proof.context -> cterm -> thm list * cterm |
21844 | 20 |
val trans_terms: Proof.context -> thm list -> thm |
21 |
val trans_props: Proof.context -> thm list -> thm |
|
35739 | 22 |
val contract: Proof.context -> thm list -> cterm -> thm -> thm |
21506 | 23 |
val print_rules: Proof.context -> unit |
18840 | 24 |
val defn_add: attribute |
25 |
val defn_del: attribute |
|
23541 | 26 |
val meta_rewrite_conv: Proof.context -> conv |
21506 | 27 |
val meta_rewrite_rule: Proof.context -> thm -> thm |
20306 | 28 |
val unfold: Proof.context -> thm list -> thm -> thm |
29 |
val unfold_goals: Proof.context -> thm list -> thm -> thm |
|
30 |
val unfold_tac: Proof.context -> thm list -> tactic |
|
31 |
val fold: Proof.context -> thm list -> thm -> thm |
|
32 |
val fold_tac: Proof.context -> thm list -> tactic |
|
33 |
val derived_def: Proof.context -> bool -> term -> |
|
20909 | 34 |
((string * typ) * term) * (Proof.context -> thm -> thm) |
18830 | 35 |
end; |
36 |
||
35624 | 37 |
structure Local_Defs: LOCAL_DEFS = |
18830 | 38 |
struct |
39 |
||
18840 | 40 |
(** primitive definitions **) |
41 |
||
18830 | 42 |
(* prepare defs *) |
43 |
||
44 |
fun cert_def ctxt eq = |
|
45 |
let |
|
40242 | 46 |
fun err msg = |
47 |
cat_error msg ("The error(s) above occurred in definition:\n" ^ |
|
48 |
quote (Syntax.string_of_term ctxt eq)); |
|
18950 | 49 |
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
|
50 |
|> Sign.no_vars ctxt |
33385 | 51 |
|> Primitive_Defs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true) |
18950 | 52 |
handle TERM (msg, _) => err msg | ERROR msg => err msg; |
53 |
in (Term.dest_Free (Term.head_of lhs), eq') end; |
|
18830 | 54 |
|
33385 | 55 |
val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free; |
18830 | 56 |
|
18875 | 57 |
fun mk_def ctxt args = |
58 |
let |
|
42501
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
59 |
val (bs, rhss) = split_list args; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
60 |
val Ts = map Term.fastype_of rhss; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
61 |
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
|
62 |
val lhss = ListPair.map Free (xs, Ts); |
18875 | 63 |
in map Logic.mk_equals (lhss ~~ rhss) end; |
64 |
||
18830 | 65 |
|
66 |
(* export defs *) |
|
67 |
||
20021 | 68 |
val head_of_def = |
20887 | 69 |
#1 o Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body; |
20021 | 70 |
|
18830 | 71 |
|
18875 | 72 |
(* |
20887 | 73 |
[x, x == a] |
18896 | 74 |
: |
75 |
B x |
|
76 |
----------- |
|
20887 | 77 |
B a |
18875 | 78 |
*) |
21684 | 79 |
fun expand defs = |
80 |
Drule.implies_intr_list defs |
|
81 |
#> Drule.generalize ([], map (head_of_def o Thm.term_of) defs) |
|
82 |
#> funpow (length defs) (fn th => Drule.reflexive_thm RS th); |
|
83 |
||
21801 | 84 |
val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of); |
21684 | 85 |
|
86 |
fun def_export _ defs = (expand defs, expand_term defs); |
|
18830 | 87 |
|
88 |
||
89 |
(* add defs *) |
|
90 |
||
20887 | 91 |
fun add_defs defs ctxt = |
18830 | 92 |
let |
42496 | 93 |
val ((xs, mxs), specs) = defs |> split_list |>> split_list; |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
26424
diff
changeset
|
94 |
val ((bfacts, atts), rhss) = specs |> split_list |>> split_list; |
42496 | 95 |
val names = map2 Thm.def_binding_optional xs bfacts; |
20887 | 96 |
val eqs = mk_def ctxt (xs ~~ rhss); |
97 |
val lhss = map (fst o Logic.dest_equals) eqs; |
|
18830 | 98 |
in |
99 |
ctxt |
|
42496 | 100 |
|> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2 |
20980 | 101 |
|> fold Variable.declare_term eqs |
42360 | 102 |
|> Proof_Context.add_assms_i def_export |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
26424
diff
changeset
|
103 |
(map2 (fn a => fn eq => (a, [(eq, [])])) (names ~~ atts) eqs) |
20887 | 104 |
|>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss |
18830 | 105 |
end; |
106 |
||
25104 | 107 |
fun add_def (var, rhs) ctxt = |
30211 | 108 |
let val ([(lhs, (_, th))], ctxt') = add_defs [(var, (Thm.empty_binding, rhs))] ctxt |
25104 | 109 |
in ((lhs, th), ctxt') end; |
25025 | 110 |
|
18840 | 111 |
|
25119 | 112 |
(* fixed_abbrev *) |
113 |
||
114 |
fun fixed_abbrev ((x, mx), rhs) ctxt = |
|
115 |
let |
|
116 |
val T = Term.fastype_of rhs; |
|
117 |
val ([x'], ctxt') = ctxt |
|
118 |
|> Variable.declare_term rhs |
|
42360 | 119 |
|> Proof_Context.add_fixes [(x, SOME T, mx)]; |
25119 | 120 |
val lhs = Free (x', T); |
121 |
val _ = cert_def ctxt' (Logic.mk_equals (lhs, rhs)); |
|
122 |
fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]); |
|
123 |
val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt'; |
|
124 |
in ((lhs, rhs), ctxt'') end; |
|
125 |
||
126 |
||
24985 | 127 |
(* specific export -- result based on educated guessing *) |
21568 | 128 |
|
24985 | 129 |
(* |
130 |
[xs, xs == as] |
|
131 |
: |
|
132 |
B xs |
|
133 |
-------------- |
|
134 |
B as |
|
135 |
*) |
|
136 |
fun export inner outer = (*beware of closure sizes*) |
|
21568 | 137 |
let |
22671 | 138 |
val exp = Assumption.export false inner outer; |
30473
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
wenzelm
parents:
30434
diff
changeset
|
139 |
val prems = Assumption.all_prems_of inner; |
22671 | 140 |
in fn th => |
141 |
let |
|
142 |
val th' = exp th; |
|
22691 | 143 |
val still_fixed = map #1 (Thm.fold_terms Term.add_frees th' []); |
22671 | 144 |
val defs = prems |> filter_out (fn prem => |
145 |
(case try (head_of_def o Thm.prop_of) prem of |
|
146 |
SOME x => member (op =) still_fixed x |
|
147 |
| NONE => true)); |
|
148 |
in (map Drule.abs_def defs, th') end |
|
149 |
end; |
|
21568 | 150 |
|
24985 | 151 |
(* |
152 |
[xs, xs == as] |
|
153 |
: |
|
154 |
TERM b xs |
|
155 |
-------------- and -------------- |
|
156 |
TERM b as b xs == b as |
|
157 |
*) |
|
158 |
fun export_cterm inner outer ct = |
|
35717 | 159 |
export inner outer (Drule.mk_term ct) ||> Drule.dest_term; |
24985 | 160 |
|
21568 | 161 |
|
21844 | 162 |
(* basic transitivity reasoning -- modulo beta-eta *) |
163 |
||
164 |
local |
|
165 |
||
166 |
val is_trivial = Pattern.aeconv o Logic.dest_equals o Thm.prop_of; |
|
167 |
||
168 |
fun trans_list _ _ [] = raise Empty |
|
169 |
| trans_list trans ctxt (th :: raw_eqs) = |
|
170 |
(case filter_out is_trivial raw_eqs of |
|
171 |
[] => th |
|
172 |
| eqs => |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30763
diff
changeset
|
173 |
let val ((_, th' :: eqs'), ctxt') = Variable.import true (th :: eqs) ctxt |
21844 | 174 |
in singleton (Variable.export ctxt' ctxt) (fold trans eqs' th') end); |
175 |
||
176 |
in |
|
177 |
||
178 |
val trans_terms = trans_list (fn eq2 => fn eq1 => eq2 COMP (eq1 COMP Drule.transitive_thm)); |
|
179 |
val trans_props = trans_list (fn eq => fn th => th COMP (eq COMP Drule.equal_elim_rule1)); |
|
180 |
||
181 |
end; |
|
182 |
||
35739 | 183 |
fun contract ctxt defs ct th = |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
184 |
trans_props ctxt [th, Thm.symmetric (Raw_Simplifier.rewrite true defs ct)]; |
21844 | 185 |
|
18840 | 186 |
|
187 |
(** defived definitions **) |
|
188 |
||
189 |
(* transformation rules *) |
|
190 |
||
33519 | 191 |
structure Rules = Generic_Data |
18840 | 192 |
( |
193 |
type T = thm list; |
|
33519 | 194 |
val empty = []; |
18840 | 195 |
val extend = I; |
33519 | 196 |
val merge = Thm.merge_thms; |
18840 | 197 |
); |
198 |
||
22846 | 199 |
fun print_rules ctxt = |
200 |
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
|
201 |
(map (Display.pretty_thm ctxt) (Rules.get (Context.Proof ctxt)))); |
18840 | 202 |
|
24039
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23541
diff
changeset
|
203 |
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
|
204 |
val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm); |
18840 | 205 |
|
206 |
||
18875 | 207 |
(* meta rewrite rules *) |
18840 | 208 |
|
23541 | 209 |
fun meta_rewrite_conv ctxt = |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
210 |
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
|
211 |
(Raw_Simplifier.context ctxt empty_ss |
26424 | 212 |
addeqcongs [Drule.equals_cong] (*protect meta-level equality*) |
213 |
addsimps (Rules.get (Context.Proof ctxt))); |
|
18840 | 214 |
|
23541 | 215 |
val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv; |
18840 | 216 |
|
18875 | 217 |
|
218 |
(* rewriting with object-level rules *) |
|
219 |
||
21506 | 220 |
fun meta f ctxt = f o map (meta_rewrite_rule ctxt); |
18840 | 221 |
|
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
222 |
val unfold = meta Raw_Simplifier.rewrite_rule; |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
223 |
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
|
224 |
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
|
225 |
val fold = meta Raw_Simplifier.fold_rule; |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
40242
diff
changeset
|
226 |
val fold_tac = meta Raw_Simplifier.fold_goals_tac; |
18840 | 227 |
|
228 |
||
229 |
(* derived defs -- potentially within the object-logic *) |
|
230 |
||
18950 | 231 |
fun derived_def ctxt conditional prop = |
18840 | 232 |
let |
233 |
val ((c, T), rhs) = prop |
|
42360 | 234 |
|> Thm.cterm_of (Proof_Context.theory_of ctxt) |
23541 | 235 |
|> meta_rewrite_conv ctxt |
18840 | 236 |
|> (snd o Logic.dest_equals o Thm.prop_of) |
21568 | 237 |
|> conditional ? Logic.strip_imp_concl |
18950 | 238 |
|> (abs_def o #2 o cert_def ctxt); |
20909 | 239 |
fun prove ctxt' def = |
42482 | 240 |
Goal.prove ctxt' (Variable.add_free_names ctxt' prop []) [] prop (K (ALLGOALS |
241 |
(CONVERSION (meta_rewrite_conv ctxt') THEN' |
|
242 |
rewrite_goal_tac [def] THEN' |
|
243 |
resolve_tac [Drule.reflexive_thm]))) |
|
244 |
handle ERROR msg => cat_error msg "Failed to prove definitional specification"; |
|
18840 | 245 |
in (((c, T), rhs), prove) end; |
246 |
||
18830 | 247 |
end; |