1(*  Title:      Pure/Isar/local_defs.ML
2    Author:     Makarius
3
4Local definitions.
5*)
6
7signature LOCAL_DEFS =
8sig
9  val cert_def: Proof.context -> (string -> Position.T list) -> term -> (string * typ) * term
10  val abs_def: term -> (string * typ) * term
11  val expand: cterm list -> thm -> thm
12  val def_export: Assumption.export
13  val define: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context ->
14    (term * (string * thm)) list * Proof.context
15  val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->
16    (term * term) * Proof.context
17  val export: Proof.context -> Proof.context -> thm -> (thm list * thm list) * thm
18  val export_cterm: Proof.context -> Proof.context -> cterm -> (thm list * thm list) * cterm
19  val contract: Proof.context -> thm list -> cterm -> thm -> thm
20  val print_rules: Proof.context -> unit
21  val defn_add: attribute
22  val defn_del: attribute
23  val meta_rewrite_conv: Proof.context -> conv
24  val meta_rewrite_rule: Proof.context -> thm -> thm
25  val abs_def_rule: Proof.context -> thm -> thm
26  val unfold_abs_def_raw: Config.raw
27  val unfold_abs_def: bool Config.T
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 unfold0: Proof.context -> thm list -> thm -> thm
32  val unfold0_goals: Proof.context -> thm list -> thm -> thm
33  val unfold0_tac: Proof.context -> thm list -> tactic
34  val fold: Proof.context -> thm list -> thm -> thm
35  val fold_tac: Proof.context -> thm list -> tactic
36  val derived_def: Proof.context -> (string -> Position.T list) -> {conditional: bool} ->
37    term -> ((string * typ) * term) * (Proof.context -> thm -> thm)
38end;
39
40structure Local_Defs: LOCAL_DEFS =
41struct
42
43(** primitive definitions **)
44
45(* prepare defs *)
46
47fun cert_def ctxt get_pos eq =
48  let
49    fun err msg =
50      cat_error msg ("The error(s) above occurred in definition:\n" ^
51        quote (Syntax.string_of_term ctxt eq));
52    val ((lhs, _), args, eq') = eq
53      |> Sign.no_vars ctxt
54      |> Primitive_Defs.dest_def ctxt
55        {check_head = Term.is_Free,
56         check_free_lhs = not o Variable.is_fixed ctxt,
57         check_free_rhs = if Variable.is_body ctxt then K true else Variable.is_fixed ctxt,
58         check_tfree = K true}
59      handle TERM (msg, _) => err msg | ERROR msg => err msg;
60    val _ =
61      Context_Position.reports ctxt
62        (maps (fn Free (x, _) => Syntax_Phases.reports_of_scope (get_pos x) | _ => []) args);
63  in (Term.dest_Free (Term.head_of lhs), eq') end;
64
65val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free;
66
67fun mk_def ctxt args =
68  let
69    val (bs, rhss) = split_list args;
70    val Ts = map Term.fastype_of rhss;
71    val (xs, _) = ctxt
72      |> Context_Position.set_visible false
73      |> Proof_Context.add_fixes (map2 (fn b => fn T => (b, SOME T, NoSyn)) bs Ts);
74    val lhss = ListPair.map Free (xs, Ts);
75  in map Logic.mk_equals (lhss ~~ rhss) end;
76
77
78(* export defs *)
79
80val head_of_def =
81  Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body;
82
83
84(*
85  [x, x \<equiv> a]
86       :
87      B x
88  -----------
89      B a
90*)
91fun expand defs =
92  Drule.implies_intr_list defs
93  #> Drule.generalize ([], map (#1 o head_of_def o Thm.term_of) defs)
94  #> funpow (length defs) (fn th => Drule.reflexive_thm RS th);
95
96val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of);
97
98fun def_export _ defs = (expand defs, expand_term defs);
99
100
101(* define *)
102
103fun define defs ctxt =
104  let
105    val ((xs, mxs), specs) = defs |> split_list |>> split_list;
106    val (bs, rhss) = specs |> split_list;
107    val eqs = mk_def ctxt (xs ~~ rhss);
108    val lhss = map (fst o Logic.dest_equals) eqs;
109  in
110    ctxt
111    |> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2
112    |> fold Variable.declare_term eqs
113    |> Proof_Context.add_assms def_export (map2 (fn b => fn eq => (b, [(eq, [])])) bs eqs)
114    |>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss
115  end;
116
117
118(* fixed_abbrev *)
119
120fun fixed_abbrev ((x, mx), rhs) ctxt =
121  let
122    val T = Term.fastype_of rhs;
123    val ([x'], ctxt') = ctxt
124      |> Variable.declare_term rhs
125      |> Proof_Context.add_fixes [(x, SOME T, mx)];
126    val lhs = Free (x', T);
127    val _ = cert_def ctxt' (K []) (Logic.mk_equals (lhs, rhs));
128    fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]);
129    val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt';
130  in ((lhs, rhs), ctxt'') end;
131
132
133(* specific export -- result based on educated guessing *)
134
135(*
136  [xs, xs \<equiv> as]
137        :
138       B xs
139  --------------
140       B as
141*)
142fun export inner outer th =
143  let
144    val defs_asms =
145      Assumption.local_assms_of inner outer
146      |> filter_out (Drule.is_sort_constraint o Thm.term_of)
147      |> map (Thm.assume #> (fn asm =>
148        (case try (head_of_def o Thm.prop_of) asm of
149          NONE => (asm, false)
150        | SOME x =>
151            let val t = Free x in
152              (case try (Assumption.export_term inner outer) t of
153                NONE => (asm, false)
154              | SOME u =>
155                  if t aconv u then (asm, false)
156                  else (Drule.abs_def (Variable.gen_all outer asm), true))
157            end)));
158  in (apply2 (map #1) (List.partition #2 defs_asms), Assumption.export false inner outer th) end;
159
160(*
161  [xs, xs \<equiv> as]
162        :
163     TERM b xs
164  --------------  and  --------------
165     TERM b as          b xs \<equiv> b as
166*)
167fun export_cterm inner outer ct =
168  export inner outer (Drule.mk_term ct) ||> Drule.dest_term;
169
170fun contract ctxt defs ct th =
171  th COMP (Raw_Simplifier.rewrite ctxt true defs ct COMP_INCR Drule.equal_elim_rule2);
172
173
174
175(** defived definitions **)
176
177(* transformation via rewrite rules *)
178
179structure Rules = Generic_Data
180(
181  type T = thm list;
182  val empty = [];
183  val extend = I;
184  val merge = Thm.merge_thms;
185);
186
187fun print_rules ctxt =
188  Pretty.writeln (Pretty.big_list "definitional rewrite rules:"
189    (map (Thm.pretty_thm_item ctxt) (Rules.get (Context.Proof ctxt))));
190
191val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm o Thm.trim_context);
192val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm);
193
194
195(* meta rewrite rules *)
196
197fun meta_rewrite_conv ctxt =
198  Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))
199    (ctxt
200      |> Raw_Simplifier.init_simpset (Rules.get (Context.Proof ctxt))
201      |> Raw_Simplifier.add_eqcong Drule.equals_cong);    (*protect meta-level equality*)
202
203val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv;
204
205fun abs_def_rule ctxt = meta_rewrite_rule ctxt #> Drule.abs_def;
206
207
208(* unfold object-level rules *)
209
210val unfold_abs_def_raw = Config.declare ("unfold_abs_def", \<^here>) (K (Config.Bool true));
211val unfold_abs_def = Config.bool unfold_abs_def_raw;
212
213local
214
215fun gen_unfold rewrite ctxt rews =
216  let val meta_rews = map (meta_rewrite_rule ctxt) rews in
217    if Config.get ctxt unfold_abs_def then
218      rewrite ctxt meta_rews #>
219      rewrite ctxt (map (perhaps (try Drule.abs_def)) meta_rews)
220    else rewrite ctxt meta_rews
221  end;
222
223val no_unfold_abs_def = Config.put unfold_abs_def false;
224
225in
226
227val unfold = gen_unfold Raw_Simplifier.rewrite_rule;
228val unfold_goals = gen_unfold Raw_Simplifier.rewrite_goals_rule;
229val unfold_tac = PRIMITIVE oo unfold_goals;
230
231val unfold0 = unfold o no_unfold_abs_def;
232val unfold0_goals = unfold_goals o no_unfold_abs_def;
233val unfold0_tac = unfold_tac o no_unfold_abs_def;
234
235end
236
237
238(* fold object-level rules *)
239
240fun fold ctxt rews = Raw_Simplifier.fold_rule ctxt (map (meta_rewrite_rule ctxt) rews);
241fun fold_tac ctxt rews = Raw_Simplifier.fold_goals_tac ctxt (map (meta_rewrite_rule ctxt) rews);
242
243
244(* derived defs -- potentially within the object-logic *)
245
246fun derived_def ctxt get_pos {conditional} prop =
247  let
248    val ((c, T), rhs) = prop
249      |> Thm.cterm_of ctxt
250      |> meta_rewrite_conv ctxt
251      |> (snd o Logic.dest_equals o Thm.prop_of)
252      |> conditional ? Logic.strip_imp_concl
253      |> (abs_def o #2 o cert_def ctxt get_pos);
254    fun prove ctxt' def =
255      Goal.prove ctxt'
256        ((not (Variable.is_body ctxt') ? Variable.add_free_names ctxt' prop) []) [] prop
257        (fn {context = ctxt'', ...} =>
258          ALLGOALS
259            (CONVERSION (meta_rewrite_conv ctxt'') THEN'
260              rewrite_goal_tac ctxt'' [def] THEN'
261              resolve_tac ctxt'' [Drule.reflexive_thm]))
262      handle ERROR msg => cat_error msg "Failed to prove definitional specification";
263  in (((c, T), rhs), prove) end;
264
265end;
266