author | wenzelm |
Mon, 15 Feb 2016 14:55:44 +0100 | |
changeset 62337 | d3996d5873dd |
parent 62137 | b8dc1fd7d900 |
child 62324 | ae44f16dcea5 |
permissions | -rw-r--r-- |
60918 | 1 |
(* Title: HOL/Tools/BNF/bnf_lift.ML |
2 |
Author: Julian Biendarra, TU Muenchen |
|
3 |
Author: Dmitriy Traytel, ETH Zurich |
|
4 |
Copyright 2015 |
|
5 |
||
6 |
Lifting of BNFs through typedefs. |
|
7 |
*) |
|
8 |
||
61067 | 9 |
signature BNF_LIFT = |
10 |
sig |
|
11 |
datatype lift_bnf_option = |
|
12 |
Plugins_Option of Proof.context -> Plugin_Name.filter |
|
13 |
| No_Warn_Wits |
|
60918 | 14 |
val copy_bnf: |
15 |
(((lift_bnf_option list * (binding option * (string * sort option)) list) * |
|
16 |
string) * thm option) * (binding * binding) -> |
|
17 |
local_theory -> local_theory |
|
18 |
val copy_bnf_cmd: |
|
19 |
(((lift_bnf_option list * (binding option * (string * string option)) list) * |
|
20 |
string) * (Facts.ref * Token.src list) option) * (binding * binding) -> |
|
21 |
local_theory -> local_theory |
|
22 |
val lift_bnf: |
|
23 |
(((lift_bnf_option list * (binding option * (string * sort option)) list) * |
|
24 |
string) * thm option) * (binding * binding) -> |
|
25 |
({context: Proof.context, prems: thm list} -> tactic) list -> |
|
26 |
local_theory -> local_theory |
|
27 |
val lift_bnf_cmd: |
|
28 |
((((lift_bnf_option list * (binding option * (string * string option)) list) * |
|
29 |
string) * string list) * (Facts.ref * Token.src list) option) * (binding * binding) -> |
|
30 |
local_theory -> Proof.state |
|
61067 | 31 |
end |
60918 | 32 |
|
61067 | 33 |
structure BNF_Lift : BNF_LIFT = |
34 |
struct |
|
60918 | 35 |
|
36 |
open Ctr_Sugar_Tactics |
|
37 |
open BNF_Util |
|
38 |
open BNF_Comp |
|
39 |
open BNF_Def |
|
40 |
||
61067 | 41 |
|
42 |
(* typedef_bnf *) |
|
43 |
||
44 |
datatype lift_bnf_option = |
|
45 |
Plugins_Option of Proof.context -> Plugin_Name.filter |
|
46 |
| No_Warn_Wits; |
|
60918 | 47 |
|
48 |
fun typedef_bnf thm wits specs map_b rel_b opts lthy = |
|
49 |
let |
|
61073 | 50 |
val plugins = |
51 |
get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) |
|
60918 | 52 |
|> the_default Plugin_Name.default_filter; |
61072 | 53 |
val no_warn_wits = exists (fn No_Warn_Wits => true | _ => false) opts; |
60918 | 54 |
|
55 |
(* extract Rep Abs F RepT AbsT *) |
|
61073 | 56 |
val (_, [Rep_G, Abs_G, F]) = Term.strip_comb (HOLogic.dest_Trueprop (Thm.prop_of thm)); |
57 |
val typ_Abs_G = dest_funT (fastype_of Abs_G); |
|
60918 | 58 |
val RepT = fst typ_Abs_G; (* F *) |
59 |
val AbsT = snd typ_Abs_G; (* G *) |
|
60 |
val AbsT_name = fst (dest_Type AbsT); |
|
61 |
val tvs = AbsT |> dest_Type |> snd |> map (fst o dest_TVar); |
|
62 |
val alpha0s = map (TFree o snd) specs; |
|
61067 | 63 |
|
60918 | 64 |
(* instantiate the new type variables newtvs to oldtvs *) |
65 |
val subst = subst_TVars (tvs ~~ alpha0s); |
|
66 |
val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); |
|
67 |
||
68 |
val Rep_G = subst Rep_G; |
|
69 |
val Abs_G = subst Abs_G; |
|
70 |
val F = subst F; |
|
71 |
val RepT = typ_subst RepT; |
|
72 |
val AbsT = typ_subst AbsT; |
|
73 |
||
61073 | 74 |
fun flatten_tyargs Ass = |
75 |
map dest_TFree alpha0s |
|
76 |
|> filter (fn T => exists (fn Ts => member (op =) Ts T) Ass); |
|
60918 | 77 |
|
78 |
val Ds0 = filter (is_none o fst) specs |> map snd; |
|
79 |
||
80 |
(* get the bnf for RepT *) |
|
81 |
val ((bnf, (deads, alphas)),((_, unfolds), lthy)) = |
|
82 |
bnf_of_typ Dont_Inline (Binding.qualify true AbsT_name) flatten_tyargs [] |
|
83 |
Ds0 RepT ((empty_comp_cache, empty_unfolds), lthy); |
|
84 |
||
61073 | 85 |
val set_bs = |
86 |
map (fn T => find_index (fn U => T = U) alpha0s) alphas |
|
60918 | 87 |
|> map (the_default Binding.empty o fst o nth specs); |
88 |
||
61067 | 89 |
val _ = (case alphas of [] => error "No live variables" | _ => ()); |
60918 | 90 |
|
91 |
val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds; |
|
92 |
||
93 |
(* number of live variables *) |
|
94 |
val lives = length alphas; |
|
95 |
||
96 |
(* state the three required properties *) |
|
97 |
val sorts = map Type.sort_of_atyp alphas; |
|
98 |
val names_lthy = fold Variable.declare_typ (alphas @ deads) lthy; |
|
99 |
val (alphas', names_lthy) = mk_TFrees' sorts names_lthy; |
|
100 |
val (betas, names_lthy) = mk_TFrees' sorts names_lthy; |
|
101 |
||
102 |
val map_F = mk_map_of_bnf deads alphas betas bnf; |
|
103 |
||
104 |
val (typ_fs, typ_aF) = fastype_of map_F |> strip_typeN lives ||> domain_type; |
|
105 |
val typ_pairs = map HOLogic.mk_prodT (alphas ~~ alphas'); |
|
106 |
val typ_subst_pair = typ_subst_atomic (alphas ~~ typ_pairs); |
|
107 |
val typ_pair = typ_subst_pair RepT; |
|
108 |
||
109 |
val subst_b = subst_atomic_types (alphas ~~ betas); |
|
110 |
val subst_a' = subst_atomic_types (alphas ~~ alphas'); |
|
111 |
val subst_pair = subst_atomic_types (alphas ~~ typ_pairs); |
|
112 |
val aF_set = F; |
|
113 |
val bF_set = subst_b F; |
|
114 |
val aF_set' = subst_a' F; |
|
115 |
val pairF_set = subst_pair F; |
|
116 |
val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf; |
|
117 |
val map_F_snd = mk_map_of_bnf deads typ_pairs alphas' bnf; |
|
118 |
val wits_F = mk_wits_of_bnf |
|
119 |
(replicate (nwits_of_bnf bnf) deads) |
|
120 |
(replicate (nwits_of_bnf bnf) alphas) bnf; |
|
121 |
||
122 |
(* val map_closed_F = @{term "\<And>f x. x \<in> F \<Longrightarrow> map_F f x \<in> F"}; *) |
|
123 |
val (var_fs, names_lthy) = mk_Frees "f" typ_fs names_lthy; |
|
124 |
val (var_x, names_lthy) = mk_Frees "x" [typ_aF] names_lthy |>> the_single; |
|
61073 | 125 |
val mem_x = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_x, aF_set)); |
60918 | 126 |
val map_f = list_comb (map_F, var_fs); |
61073 | 127 |
val mem_map = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_f $ var_x, bF_set)); |
60918 | 128 |
val imp_map = Logic.mk_implies (mem_x, mem_map); |
61073 | 129 |
val map_closed_F = fold_rev Logic.all var_fs (Logic.all var_x imp_map); |
60918 | 130 |
|
131 |
(* val zip_closed_F = @{term "\<And>z. map_F fst z \<in> F \<Longrightarrow> map_F snd z \<in> F \<Longrightarrow> z \<in> F"}; *) |
|
132 |
val (var_zs, names_lthy) = mk_Frees "z" [typ_pair] names_lthy; |
|
133 |
val (pairs, names_lthy) = mk_Frees "tmp" typ_pairs names_lthy; |
|
134 |
val var_z = hd var_zs; |
|
135 |
val fsts = map (fst o Term.strip_comb o HOLogic.mk_fst) pairs; |
|
136 |
val snds = map (fst o Term.strip_comb o HOLogic.mk_snd) pairs; |
|
137 |
val map_fst = list_comb (list_comb (map_F_fst, fsts), var_zs); |
|
61073 | 138 |
val mem_map_fst = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_fst, aF_set)); |
60918 | 139 |
val map_snd = list_comb (list_comb (map_F_snd, snds), var_zs); |
61073 | 140 |
val mem_map_snd = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_snd, aF_set')); |
141 |
val mem_z = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_z, pairF_set)); |
|
60918 | 142 |
val imp_zip = Logic.mk_implies (mem_map_fst, Logic.mk_implies (mem_map_snd, mem_z)); |
143 |
val zip_closed_F = Logic.all var_z imp_zip; |
|
144 |
||
145 |
(* val wit_closed_F = @{term "wit_F a \<in> F"}; *) |
|
146 |
val (var_as, names_lthy) = mk_Frees "a" alphas names_lthy; |
|
147 |
val (var_bs, _) = mk_Frees "a" alphas names_lthy; |
|
62137 | 148 |
fun binder_types_until_eq V T = |
149 |
let |
|
150 |
fun strip (TU as Type ("fun", [T, U])) = if V = TU then [] else T :: strip U |
|
151 |
| strip T = if V = T then [] else |
|
152 |
error ("Bad type for witness: " ^ quote (Syntax.string_of_typ lthy T)); |
|
153 |
in strip T end; |
|
60918 | 154 |
val Iwits = the_default wits_F (Option.map (map (`(map (fn T => |
62137 | 155 |
find_index (fn U => T = U) alphas) o binder_types_until_eq RepT o fastype_of))) wits); |
60918 | 156 |
val wit_closed_Fs = |
61067 | 157 |
Iwits |> map (fn (I, wit_F) => |
60918 | 158 |
let |
159 |
val vars = map (nth var_as) I; |
|
160 |
val wit_a = list_comb (wit_F, vars); |
|
61073 | 161 |
in fold_rev Logic.all vars (HOLogic.mk_Trueprop (HOLogic.mk_mem (wit_a, aF_set))) end); |
60918 | 162 |
|
163 |
val mk_wit_goals = mk_wit_goals var_as var_bs |
|
164 |
(mk_sets_of_bnf (replicate lives deads) (replicate lives alphas) bnf); |
|
165 |
||
166 |
val goals = [map_closed_F, zip_closed_F] @ wit_closed_Fs @ |
|
167 |
(case wits of NONE => [] | _ => maps mk_wit_goals Iwits); |
|
168 |
||
169 |
val lost_wits = filter_out (fn (J, _) => exists (fn (I, _) => I = J) Iwits) wits_F; |
|
61067 | 170 |
val _ = |
171 |
if null lost_wits orelse no_warn_wits then () |
|
172 |
else |
|
173 |
lost_wits |
|
174 |
|> map (Syntax.pretty_typ lthy o fastype_of o snd) |
|
175 |
|> Pretty.big_list |
|
176 |
"The following types of nonemptiness witnesses of the raw type's BNF were lost:" |
|
177 |
|> (fn pt => Pretty.chunks [pt, |
|
178 |
Pretty.para "You can specify a liftable witness (e.g., a term of one of the above types\ |
|
179 |
\ that satisfies the typedef's invariant)\ |
|
180 |
\ using the annotation [wits: <term>]."]) |
|
181 |
|> Pretty.string_of |
|
182 |
|> warning; |
|
60918 | 183 |
|
184 |
fun after_qed ([map_closed_thm] :: [zip_closed_thm] :: wit_thmss) lthy = |
|
61067 | 185 |
let |
186 |
val (wit_closed_thms, wit_thms) = |
|
187 |
(case wits of |
|
188 |
NONE => (map the_single wit_thmss, wit_thms_of_bnf bnf) |
|
189 |
| _ => chop (length wit_closed_Fs) (map the_single wit_thmss)) |
|
60918 | 190 |
|
61067 | 191 |
(* construct map set bd rel wit *) |
192 |
(* val map_G = @{term "\<lambda>f. Abs_G o map_F f o Rep_G"}; *) |
|
193 |
val Abs_Gb = subst_b Abs_G; |
|
61073 | 194 |
val map_G = |
195 |
fold_rev HOLogic.tupled_lambda var_fs |
|
196 |
(HOLogic.mk_comp (HOLogic.mk_comp (Abs_Gb, map_f), Rep_G)); |
|
60918 | 197 |
|
61067 | 198 |
(* val sets_G = [@{term "set_F o Rep_G"}]; *) |
199 |
val sets_F = mk_sets_of_bnf (replicate lives deads) (replicate lives alphas) bnf; |
|
200 |
val sets_G = map (fn set_F => HOLogic.mk_comp (set_F, Rep_G)) sets_F; |
|
60918 | 201 |
|
61067 | 202 |
(* val bd_G = @{term "bd_F"}; *) |
203 |
val bd_F = mk_bd_of_bnf deads alphas bnf; |
|
204 |
val bd_G = bd_F; |
|
205 |
||
206 |
(* val rel_G = @{term "\<lambda>R. BNF_Def.vimage2p Rep_G Rep_G (rel_F R)"}; *) |
|
207 |
val rel_F = mk_rel_of_bnf deads alphas betas bnf; |
|
61073 | 208 |
val (typ_Rs, _) = strip_typeN lives (fastype_of rel_F); |
60918 | 209 |
|
61067 | 210 |
val (var_Rs, names_lthy) = mk_Frees "R" typ_Rs lthy; |
211 |
val Rep_Gb = subst_b Rep_G; |
|
212 |
val rel_G = fold_rev absfree (map dest_Free var_Rs) |
|
213 |
(mk_vimage2p Rep_G Rep_Gb $ list_comb (rel_F, var_Rs)); |
|
60918 | 214 |
|
61067 | 215 |
(* val wits_G = [@{term "Abs_G o wit_F"}]; *) |
216 |
val (var_as, _) = mk_Frees "a" alphas names_lthy; |
|
217 |
val wits_G = |
|
218 |
map (fn (I, wit_F) => |
|
219 |
let |
|
220 |
val vs = map (nth var_as) I; |
|
221 |
in fold_rev absfree (map dest_Free vs) (Abs_G $ (list_comb (wit_F, vs))) end) |
|
222 |
Iwits; |
|
60918 | 223 |
|
61067 | 224 |
(* tactics *) |
225 |
val Rep_thm = thm RS @{thm type_definition.Rep}; |
|
226 |
val Abs_inverse_thm = thm RS @{thm type_definition.Abs_inverse}; |
|
227 |
val Abs_inject_thm = thm RS @{thm type_definition.Abs_inject}; |
|
228 |
val Rep_cases_thm = thm RS @{thm type_definition.Rep_cases}; |
|
229 |
val Rep_inverse_thm = thm RS @{thm type_definition.Rep_inverse}; |
|
60918 | 230 |
|
61067 | 231 |
fun map_id0_tac ctxt = |
232 |
HEADGOAL (EVERY' [rtac ctxt ext, |
|
233 |
SELECT_GOAL (unfold_thms_tac ctxt [map_id0_of_bnf bnf, id_apply, o_apply, |
|
234 |
Rep_inverse_thm]), |
|
235 |
rtac ctxt refl]); |
|
60918 | 236 |
|
61067 | 237 |
fun map_comp0_tac ctxt = |
60918 | 238 |
HEADGOAL (EVERY' [rtac ctxt ext, |
61067 | 239 |
SELECT_GOAL (unfold_thms_tac ctxt [map_comp0_of_bnf bnf, o_apply, |
60918 | 240 |
Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), |
61067 | 241 |
rtac ctxt refl]); |
60918 | 242 |
|
61067 | 243 |
fun map_cong0_tac ctxt = |
244 |
HEADGOAL (EVERY' ([SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), |
|
245 |
rtac ctxt (([Rep_thm RS map_closed_thm, Rep_thm RS map_closed_thm] MRS |
|
246 |
Abs_inject_thm) RS iffD2), |
|
247 |
rtac ctxt (map_cong0_of_bnf bnf)] @ replicate lives (Goal.assume_rule_tac ctxt))); |
|
60918 | 248 |
|
61067 | 249 |
val set_map0s_tac = |
250 |
map (fn set_map => fn ctxt => |
|
251 |
HEADGOAL (EVERY' [rtac ctxt ext, |
|
252 |
SELECT_GOAL (unfold_thms_tac ctxt [set_map, o_apply, |
|
253 |
Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), |
|
254 |
rtac ctxt refl])) |
|
255 |
(set_map_of_bnf bnf); |
|
256 |
||
257 |
fun card_order_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_card_order_of_bnf bnf)); |
|
60918 | 258 |
|
61067 | 259 |
fun cinfinite_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_cinfinite_of_bnf bnf)); |
260 |
||
261 |
val set_bds_tac = |
|
262 |
map (fn set_bd => fn ctxt => |
|
263 |
HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt set_bd])) |
|
264 |
(set_bd_of_bnf bnf); |
|
265 |
||
266 |
fun le_rel_OO_tac ctxt = |
|
267 |
HEADGOAL (EVERY' [rtac ctxt @{thm vimage2p_relcompp_mono}, |
|
268 |
rtac ctxt ((rel_OO_of_bnf bnf RS sym) RS @{thm ord_eq_le_trans}), |
|
269 |
rtac ctxt @{thm order_refl}]); |
|
60918 | 270 |
|
61067 | 271 |
fun rel_OO_Grp_tac ctxt = |
272 |
HEADGOAL (EVERY' ([SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt ext))), |
|
273 |
SELECT_GOAL (unfold_thms_tac ctxt [@{thm OO_Grp_alt}, mem_Collect_eq, |
|
274 |
o_apply, @{thm vimage2p_def}, in_rel_of_bnf bnf, Bex_def, mem_Collect_eq]), |
|
275 |
rtac ctxt iffI, |
|
276 |
SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve0_tac [exE,conjE]))), |
|
277 |
rtac ctxt (zip_closed_thm OF (replicate 2 (Rep_thm RSN (2, @{thm ssubst_mem}))) RS |
|
278 |
Rep_cases_thm), |
|
279 |
assume_tac ctxt, |
|
280 |
assume_tac ctxt, |
|
281 |
hyp_subst_tac ctxt, |
|
282 |
SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt exI))), |
|
283 |
rtac ctxt conjI] @ |
|
284 |
replicate (lives - 1) (rtac ctxt conjI THEN' assume_tac ctxt) @ |
|
285 |
[assume_tac ctxt, |
|
286 |
SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt conjI))), |
|
287 |
REPEAT_DETERM_N 2 o |
|
288 |
etac ctxt (trans OF [iffD2 OF [Abs_inject_thm OF |
|
289 |
[map_closed_thm OF [Rep_thm], Rep_thm]], Rep_inverse_thm]), |
|
290 |
SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve0_tac [exE,conjE]))), |
|
291 |
rtac ctxt exI, |
|
292 |
rtac ctxt conjI] @ |
|
293 |
replicate (lives - 1) (rtac ctxt conjI THEN' assume_tac ctxt) @ |
|
294 |
[assume_tac ctxt, |
|
295 |
rtac ctxt conjI, |
|
296 |
REPEAT_DETERM_N 2 o EVERY' |
|
297 |
[rtac ctxt (iffD1 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]]), |
|
298 |
etac ctxt (Rep_inverse_thm RS sym RSN (2, trans))]])); |
|
60918 | 299 |
|
61067 | 300 |
fun wit_tac ctxt = |
301 |
HEADGOAL (EVERY' |
|
302 |
(map (fn thm => (EVERY' |
|
303 |
[SELECT_GOAL (unfold_thms_tac ctxt (o_apply :: |
|
304 |
(wit_closed_thms RL [Abs_inverse_thm]))), |
|
305 |
dtac ctxt thm, assume_tac ctxt])) |
|
306 |
wit_thms)); |
|
60918 | 307 |
|
61067 | 308 |
val tactics = [map_id0_tac, map_comp0_tac, map_cong0_tac] @ set_map0s_tac @ |
309 |
[card_order_bd_tac, cinfinite_bd_tac] @ set_bds_tac @ [le_rel_OO_tac, rel_OO_Grp_tac]; |
|
60918 | 310 |
|
61067 | 311 |
val (bnf, lthy) = bnf_def Dont_Inline (user_policy Note_Some) false I |
312 |
tactics wit_tac NONE map_b rel_b set_bs |
|
313 |
((((((Binding.empty, AbsT), map_G), sets_G), bd_G), wits_G), SOME rel_G) |
|
314 |
lthy; |
|
60928
141a1d485259
unfold intermediate definitions (stemming from composition) in lifted bnf operations
traytel
parents:
60918
diff
changeset
|
315 |
|
61067 | 316 |
val bnf = morph_bnf_defs (Morphism.thm_morphism "BNF" (unfold_thms lthy defs)) bnf; |
317 |
in |
|
318 |
lthy |> BNF_Def.register_bnf plugins AbsT_name bnf |
|
319 |
end |
|
320 |
| after_qed _ _ = raise Match; |
|
60918 | 321 |
in |
322 |
(goals, after_qed, defs, lthy) |
|
323 |
end; |
|
324 |
||
61067 | 325 |
|
326 |
(* main commands *) |
|
327 |
||
328 |
local |
|
329 |
||
60918 | 330 |
fun prepare_common prepare_name prepare_sort prepare_term prepare_thm |
331 |
(((((plugins, raw_specs), raw_Tname), raw_wits), xthm_opt), (map_b, rel_b)) lthy = |
|
332 |
let |
|
333 |
val Tname = prepare_name lthy raw_Tname; |
|
334 |
val input_thm = |
|
335 |
(case xthm_opt of |
|
336 |
SOME xthm => prepare_thm lthy xthm |
|
337 |
| NONE => Typedef.get_info lthy Tname |> hd |> snd |> #type_definition); |
|
61067 | 338 |
val wits = (Option.map o map) (prepare_term lthy) raw_wits; |
339 |
val specs = |
|
340 |
map (apsnd (apsnd (the_default @{sort type} o Option.map (prepare_sort lthy)))) raw_specs; |
|
60918 | 341 |
|
61067 | 342 |
val _ = |
343 |
(case HOLogic.dest_Trueprop (Thm.prop_of input_thm) of |
|
344 |
Const (@{const_name type_definition}, _) $ _ $ _ $ _ => () |
|
345 |
| _ => error "Unsupported type of a theorem: only type_definition is supported"); |
|
60918 | 346 |
in |
347 |
typedef_bnf input_thm wits specs map_b rel_b plugins lthy |
|
348 |
end; |
|
349 |
||
350 |
fun prepare_lift_bnf prepare_name prepare_sort prepare_term prepare_thm = |
|
351 |
(fn (goals, after_qed, definitions, lthy) => |
|
352 |
lthy |
|
353 |
|> Proof.theorem NONE after_qed (map (single o rpair []) goals) |
|
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61073
diff
changeset
|
354 |
|> Proof.refine_singleton |
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61073
diff
changeset
|
355 |
(Method.Basic (fn ctxt => SIMPLE_METHOD (unfold_thms_tac ctxt definitions))) |
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61073
diff
changeset
|
356 |
|> Proof.refine_singleton (Method.primitive_text (K I))) oo |
60918 | 357 |
prepare_common prepare_name prepare_sort prepare_term prepare_thm o apfst (apfst (apsnd SOME)); |
358 |
||
359 |
fun prepare_solve prepare_name prepare_typ prepare_sort prepare_thm tacs = |
|
360 |
(fn (goals, after_qed, _, lthy) => |
|
361 |
lthy |
|
362 |
|> after_qed (map2 (single oo Goal.prove lthy [] []) goals (tacs (length goals)))) oo |
|
363 |
prepare_common prepare_name prepare_typ prepare_sort prepare_thm o apfst (apfst (rpair NONE)); |
|
364 |
||
61067 | 365 |
in |
366 |
||
367 |
val lift_bnf_cmd = |
|
368 |
prepare_lift_bnf |
|
369 |
(fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) |
|
370 |
Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms); |
|
371 |
||
60918 | 372 |
fun lift_bnf args tacs = prepare_solve (K I) (K I) (K I) (K I) (K tacs) args; |
61067 | 373 |
|
374 |
val copy_bnf = |
|
375 |
prepare_solve (K I) (K I) (K I) (K I) |
|
376 |
(fn n => replicate n (fn {context = ctxt, prems = _} => rtac ctxt UNIV_I 1)); |
|
377 |
||
378 |
val copy_bnf_cmd = |
|
379 |
prepare_solve |
|
380 |
(fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) |
|
381 |
Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms) |
|
382 |
(fn n => replicate n (fn {context = ctxt, prems = _} => rtac ctxt UNIV_I 1)); |
|
383 |
||
384 |
end; |
|
385 |
||
386 |
||
387 |
(* outer syntax *) |
|
388 |
||
389 |
local |
|
60918 | 390 |
|
391 |
val parse_wits = |
|
392 |
@{keyword "["} |-- (Parse.name --| @{keyword ":"} -- Scan.repeat Parse.term >> |
|
393 |
(fn ("wits", Ts) => Ts |
|
394 |
| (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --| |
|
395 |
@{keyword "]"} || Scan.succeed []; |
|
396 |
||
397 |
val parse_options = |
|
398 |
Scan.optional (@{keyword "("} |-- |
|
399 |
Parse.list1 (Parse.group (K "option") |
|
400 |
(Plugin_Name.parse_filter >> Plugins_Option |
|
401 |
|| Parse.reserved "no_warn_wits" >> K No_Warn_Wits)) |
|
402 |
--| @{keyword ")"}) []; |
|
403 |
||
404 |
val parse_plugins = |
|
405 |
Scan.optional (@{keyword "("} |-- Plugin_Name.parse_filter --| @{keyword ")"}) |
|
406 |
(K Plugin_Name.default_filter) >> Plugins_Option >> single; |
|
407 |
||
408 |
val parse_typedef_thm = Scan.option (Parse.reserved "via" |-- Parse.xthm); |
|
409 |
||
61067 | 410 |
in |
411 |
||
60918 | 412 |
val _ = |
413 |
Outer_Syntax.local_theory_to_proof @{command_keyword lift_bnf} |
|
414 |
"register a subtype of a bounded natural functor (BNF) as a BNF" |
|
415 |
((parse_options -- parse_type_args_named_constrained -- Parse.type_const -- parse_wits -- |
|
416 |
parse_typedef_thm -- parse_map_rel_bindings) >> lift_bnf_cmd); |
|
417 |
||
418 |
val _ = |
|
419 |
Outer_Syntax.local_theory @{command_keyword copy_bnf} |
|
420 |
"register a type copy of a bounded natural functor (BNF) as a BNF" |
|
421 |
((parse_plugins -- parse_type_args_named_constrained -- Parse.type_const -- |
|
422 |
parse_typedef_thm -- parse_map_rel_bindings) >> copy_bnf_cmd); |
|
423 |
||
61067 | 424 |
end; |
425 |
||
426 |
end; |