author | wenzelm |
Sun, 04 Aug 2024 13:24:54 +0200 | |
changeset 80634 | a90ab1ea6458 |
parent 74381 | 79f484b0e35b |
permissions | -rw-r--r-- |
62691
9bfcbab7cd99
put all 'bnf_*.ML' files together, irrespective of bootstrapping/dependency constraints
blanchet
parents:
60757
diff
changeset
|
1 |
(* Title: HOL/Tools/BNF/bnf_lfp_countable.ML |
58160 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
Copyright 2014 |
|
4 |
||
5 |
Countability tactic for BNF datatypes. |
|
6 |
*) |
|
7 |
||
8 |
signature BNF_LFP_COUNTABLE = |
|
9 |
sig |
|
58165
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
10 |
val derive_encode_injectives_thms: Proof.context -> string list -> thm list |
58161 | 11 |
val countable_datatype_tac: Proof.context -> tactic |
58160 | 12 |
end; |
13 |
||
14 |
structure BNF_LFP_Countable : BNF_LFP_COUNTABLE = |
|
15 |
struct |
|
16 |
||
17 |
open BNF_FP_Rec_Sugar_Util |
|
18 |
open BNF_Def |
|
19 |
open BNF_Util |
|
20 |
open BNF_Tactics |
|
21 |
open BNF_FP_Util |
|
22 |
open BNF_FP_Def_Sugar |
|
23 |
||
69593 | 24 |
val countableS = \<^sort>\<open>countable\<close>; |
58174 | 25 |
|
60728 | 26 |
fun nchotomy_tac ctxt nchotomy = |
60757 | 27 |
HEADGOAL (resolve_tac ctxt [nchotomy RS @{thm all_reg[rotated]}] THEN' |
28 |
REPEAT_ALL_NEW (resolve_tac ctxt [allI, impI] ORELSE' eresolve_tac ctxt [exE, disjE])); |
|
58160 | 29 |
|
64705
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
blanchet
parents:
64627
diff
changeset
|
30 |
fun meta_spec_mp_tac _ 0 = K all_tac |
60728 | 31 |
| meta_spec_mp_tac ctxt depth = |
60752 | 32 |
dtac ctxt meta_spec THEN' meta_spec_mp_tac ctxt (depth - 1) THEN' |
33 |
dtac ctxt meta_mp THEN' assume_tac ctxt; |
|
58160 | 34 |
|
60728 | 35 |
fun use_induction_hypothesis_tac ctxt = |
58161 | 36 |
DEEPEN (1, 64 (* large number *)) |
60752 | 37 |
(fn depth => meta_spec_mp_tac ctxt depth THEN' etac ctxt allE THEN' etac ctxt impE THEN' |
38 |
assume_tac ctxt THEN' assume_tac ctxt) 0; |
|
58160 | 39 |
|
58172 | 40 |
val same_ctr_simps = @{thms sum_encode_eq prod_encode_eq sum.inject prod.inject to_nat_split |
41 |
id_apply snd_conv simp_thms}; |
|
58166 | 42 |
val distinct_ctrs_simps = @{thms sum_encode_eq sum.inject sum.distinct simp_thms}; |
58160 | 43 |
|
58166 | 44 |
fun same_ctr_tac ctxt injects recs map_congs' inj_map_strongs' = |
45 |
HEADGOAL (asm_full_simp_tac |
|
46 |
(ss_only (injects @ recs @ map_congs' @ same_ctr_simps) ctxt) THEN_MAYBE' |
|
60728 | 47 |
TRY o REPEAT_ALL_NEW (rtac ctxt conjI) THEN_ALL_NEW |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
48 |
REPEAT_ALL_NEW (eresolve_tac ctxt (conjE :: inj_map_strongs')) THEN_ALL_NEW |
60752 | 49 |
(assume_tac ctxt ORELSE' use_induction_hypothesis_tac ctxt)); |
58160 | 50 |
|
51 |
fun distinct_ctrs_tac ctxt recs = |
|
58166 | 52 |
HEADGOAL (asm_full_simp_tac (ss_only (recs @ distinct_ctrs_simps) ctxt)); |
58160 | 53 |
|
54 |
fun mk_encode_injective_tac ctxt n nchotomy injects recs map_comps' inj_map_strongs' = |
|
55 |
let val ks = 1 upto n in |
|
60728 | 56 |
EVERY (maps (fn k => nchotomy_tac ctxt nchotomy :: map (fn k' => |
58160 | 57 |
if k = k' then same_ctr_tac ctxt injects recs map_comps' inj_map_strongs' |
58 |
else distinct_ctrs_tac ctxt recs) ks) ks) |
|
59 |
end; |
|
60 |
||
61 |
fun mk_encode_injectives_tac ctxt ns induct nchotomys injectss recss map_comps' inj_map_strongs' = |
|
60728 | 62 |
HEADGOAL (rtac ctxt induct) THEN |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58461
diff
changeset
|
63 |
EVERY (@{map 4} (fn n => fn nchotomy => fn injects => fn recs => |
58160 | 64 |
mk_encode_injective_tac ctxt n nchotomy injects recs map_comps' inj_map_strongs') |
65 |
ns nchotomys injectss recss); |
|
66 |
||
67 |
fun endgame_tac ctxt encode_injectives = |
|
68 |
unfold_thms_tac ctxt @{thms inj_on_def ball_UNIV} THEN |
|
60728 | 69 |
ALLGOALS (rtac ctxt exI THEN' rtac ctxt allI THEN' resolve_tac ctxt encode_injectives); |
58160 | 70 |
|
71 |
fun encode_sumN n k t = |
|
72 |
Balanced_Tree.access {init = t, |
|
74381 | 73 |
left = fn t => \<^Const>\<open>sum_encode for \<^Const>\<open>Inl \<^Type>\<open>nat\<close> \<^Type>\<open>nat\<close> for t\<close>\<close>, |
74 |
right = fn t => \<^Const>\<open>sum_encode for \<^Const>\<open>Inr \<^Type>\<open>nat\<close> \<^Type>\<open>nat\<close> for t\<close>\<close>} |
|
58160 | 75 |
n k; |
76 |
||
74381 | 77 |
fun encode_tuple [] = \<^Const>\<open>zero_class.zero \<^Type>\<open>nat\<close>\<close> |
58160 | 78 |
| encode_tuple ts = |
74381 | 79 |
Balanced_Tree.make (fn (t, u) => \<^Const>\<open>prod_encode for \<^Const>\<open>Pair \<^Type>\<open>nat\<close> \<^Type>\<open>nat\<close> for u t\<close>\<close>) ts; |
58160 | 80 |
|
81 |
fun mk_encode_funs ctxt fpTs ns ctrss0 recs0 = |
|
82 |
let |
|
83 |
val thy = Proof_Context.theory_of ctxt; |
|
84 |
||
58174 | 85 |
fun check_countable T = |
86 |
Sign.of_sort thy (T, countableS) orelse |
|
87 |
raise TYPE ("Type is not of sort " ^ Syntax.string_of_sort ctxt countableS, [T], []); |
|
88 |
||
89 |
fun mk_to_nat_checked T = |
|
69593 | 90 |
Const (\<^const_name>\<open>to_nat\<close>, tap check_countable T --> HOLogic.natT); |
58174 | 91 |
|
58160 | 92 |
val nn = length ns; |
58221 | 93 |
val recs as rec1 :: _ = map2 (mk_co_rec thy Least_FP (replicate nn HOLogic.natT)) fpTs recs0; |
58160 | 94 |
val arg_Ts = binder_fun_types (fastype_of rec1); |
95 |
val arg_Tss = Library.unflat ctrss0 arg_Ts; |
|
96 |
||
69593 | 97 |
fun mk_U (Type (\<^type_name>\<open>prod\<close>, [T1, T2])) = |
58160 | 98 |
if member (op =) fpTs T1 then T2 else HOLogic.mk_prodT (mk_U T1, mk_U T2) |
99 |
| mk_U (Type (s, Ts)) = Type (s, map mk_U Ts) |
|
100 |
| mk_U T = T; |
|
101 |
||
102 |
fun mk_nat (j, T) = |
|
103 |
if T = HOLogic.natT then |
|
104 |
SOME (Bound j) |
|
105 |
else if member (op =) fpTs T then |
|
106 |
NONE |
|
107 |
else if exists_subtype_in fpTs T then |
|
108 |
let val U = mk_U T in |
|
64627
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
blanchet
parents:
63859
diff
changeset
|
109 |
SOME (mk_to_nat_checked U $ (build_map ctxt [] [] (snd_const o fst) (T, U) $ Bound j)) |
58160 | 110 |
end |
111 |
else |
|
58174 | 112 |
SOME (mk_to_nat_checked T $ Bound j); |
58160 | 113 |
|
114 |
fun mk_arg n (k, arg_T) = |
|
115 |
let |
|
116 |
val bound_Ts = rev (binder_types arg_T); |
|
117 |
val nats = map_filter mk_nat (tag_list 0 bound_Ts); |
|
118 |
in |
|
119 |
fold (fn T => fn t => Abs (Name.uu, T, t)) bound_Ts (encode_sumN n k (encode_tuple nats)) |
|
120 |
end; |
|
121 |
||
122 |
val argss = map2 (map o mk_arg) ns (map (tag_list 1) arg_Tss); |
|
123 |
in |
|
124 |
map (fn recx => Term.list_comb (recx, flat argss)) recs |
|
125 |
end; |
|
126 |
||
58165
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
127 |
fun derive_encode_injectives_thms _ [] = [] |
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
128 |
| derive_encode_injectives_thms ctxt fpT_names0 = |
58160 | 129 |
let |
64705
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
blanchet
parents:
64627
diff
changeset
|
130 |
fun not_datatype_name s = |
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
blanchet
parents:
64627
diff
changeset
|
131 |
error (quote s ^ " is not a datatype"); |
58315 | 132 |
fun not_mutually_recursive ss = error (commas ss ^ " are not mutually recursive datatypes"); |
58160 | 133 |
|
134 |
fun lfp_sugar_of s = |
|
135 |
(case fp_sugar_of ctxt s of |
|
63859
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
blanchet
parents:
62691
diff
changeset
|
136 |
SOME (fp_sugar as {fp = Least_FP, fp_co_induct_sugar = SOME _, ...}) => fp_sugar |
64705
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
blanchet
parents:
64627
diff
changeset
|
137 |
| _ => not_datatype_name s); |
58160 | 138 |
|
58232
7b70a2b4ec9b
made new countable tactic work with sorts other than 'type'
blanchet
parents:
58221
diff
changeset
|
139 |
val fpTs0 as Type (_, var_As) :: _ = |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
74381
diff
changeset
|
140 |
map (#T o lfp_sugar_of o dest_Type_name) (#Ts (#fp_res (lfp_sugar_of (hd fpT_names0)))); |
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
74381
diff
changeset
|
141 |
val fpT_names = map dest_Type_name fpTs0; |
58160 | 142 |
|
143 |
val (As_names, _) = Variable.variant_fixes (map (fn TVar ((s, _), _) => s) var_As) ctxt; |
|
144 |
val As = |
|
58174 | 145 |
map2 (fn s => fn TVar (_, S) => TFree (s, union (op =) countableS S)) |
58160 | 146 |
As_names var_As; |
147 |
val fpTs = map (fn s => Type (s, As)) fpT_names; |
|
148 |
||
149 |
val _ = subset (op =) (fpT_names0, fpT_names) orelse not_mutually_recursive fpT_names0; |
|
150 |
||
151 |
fun mk_conjunct fpT x encode_fun = |
|
152 |
HOLogic.all_const fpT $ Abs (Name.uu, fpT, |
|
153 |
HOLogic.mk_imp (HOLogic.mk_eq (encode_fun $ x, encode_fun $ Bound 0), |
|
154 |
HOLogic.eq_const fpT $ x $ Bound 0)); |
|
155 |
||
58461 | 156 |
val fp_sugars as |
63859
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
blanchet
parents:
62691
diff
changeset
|
157 |
{fp_nesting_bnfs, fp_co_induct_sugar = SOME {common_co_inducts = induct :: _, ...}, |
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
blanchet
parents:
62691
diff
changeset
|
158 |
...} :: _ = |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
74381
diff
changeset
|
159 |
map (the o fp_sugar_of ctxt o dest_Type_name) fpTs0; |
58460 | 160 |
val ctr_sugars = map (#ctr_sugar o #fp_ctr_sugar) fp_sugars; |
58160 | 161 |
|
162 |
val ctrss0 = map #ctrs ctr_sugars; |
|
163 |
val ns = map length ctrss0; |
|
63859
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
blanchet
parents:
62691
diff
changeset
|
164 |
val recs0 = map (#co_rec o the o #fp_co_induct_sugar) fp_sugars; |
58160 | 165 |
val nchotomys = map #nchotomy ctr_sugars; |
166 |
val injectss = map #injects ctr_sugars; |
|
63859
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
blanchet
parents:
62691
diff
changeset
|
167 |
val rec_thmss = map (#co_rec_thms o the o #fp_co_induct_sugar) fp_sugars; |
58160 | 168 |
val map_comps' = map (unfold_thms ctxt @{thms comp_def} o map_comp_of_bnf) fp_nesting_bnfs; |
169 |
val inj_map_strongs' = map (Thm.permute_prems 0 ~1 o inj_map_strong_of_bnf) fp_nesting_bnfs; |
|
170 |
||
171 |
val (xs, names_ctxt) = ctxt |> mk_Frees "x" fpTs; |
|
172 |
||
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58461
diff
changeset
|
173 |
val conjuncts = @{map 3} mk_conjunct fpTs xs (mk_encode_funs ctxt fpTs ns ctrss0 recs0); |
58170
d84bab7ed89e
fixed tactic for n-way mutual recursion, n >= 4 (balanced conjunctions confuse the tactic)
blanchet
parents:
58168
diff
changeset
|
174 |
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj conjuncts); |
58160 | 175 |
in |
58296 | 176 |
Goal.prove (*no sorry*) ctxt [] [] goal (fn {context = ctxt, prems = _} => |
58160 | 177 |
mk_encode_injectives_tac ctxt ns induct nchotomys injectss rec_thmss map_comps' |
58161 | 178 |
inj_map_strongs') |
59970 | 179 |
|> HOLogic.conj_elims ctxt |
58160 | 180 |
|> Proof_Context.export names_ctxt ctxt |
70494 | 181 |
|> map (Thm.close_derivation \<^here>) |
58160 | 182 |
end; |
183 |
||
74381 | 184 |
fun get_countable_goal_type_name (\<^Const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>Ex\<close>, _) |
69593 | 185 |
$ Abs (_, Type (_, [Type (s, _), _]), Const (\<^const_name>\<open>inj_on\<close>, _) $ Bound 0 |
186 |
$ Const (\<^const_name>\<open>top\<close>, _)))) = s |
|
58165
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
187 |
| get_countable_goal_type_name _ = error "Wrong goal format for datatype countability tactic"; |
58160 | 188 |
|
58161 | 189 |
fun core_countable_datatype_tac ctxt st = |
58165
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
190 |
let val T_names = map get_countable_goal_type_name (Thm.prems_of st) in |
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
191 |
endgame_tac ctxt (derive_encode_injectives_thms ctxt T_names) st |
2ec97d9c1e83
'prove_sorry' is too dangerous here -- the tactic is sometimes applied to non-theorems
blanchet
parents:
58161
diff
changeset
|
192 |
end; |
58160 | 193 |
|
58161 | 194 |
fun countable_datatype_tac ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
195 |
TRY (Class.intro_classes_tac ctxt []) THEN core_countable_datatype_tac ctxt; |
58160 | 196 |
|
197 |
end; |