author | wenzelm |
Sat, 30 Nov 2024 13:40:57 +0100 | |
changeset 81511 | 8cbc8bc6f382 |
parent 80634 | a90ab1ea6458 |
permissions | -rw-r--r-- |
55061 | 1 |
(* Title: HOL/Tools/BNF/bnf_comp.ML |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Dmitriy Traytel, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
4 |
Copyright 2012 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
5 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
Composition of bounded natural functors. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
7 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
9 |
signature BNF_COMP = |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
10 |
sig |
59994 | 11 |
val typedef_threshold: int Config.T |
63800 | 12 |
val with_typedef_threshold: int -> (Proof.context -> Proof.context) -> Proof.context -> |
13 |
Proof.context |
|
14 |
val with_typedef_threshold_yield: int -> (Proof.context -> 'a * Proof.context) -> Proof.context -> |
|
15 |
'a * Proof.context |
|
59710 | 16 |
|
51837
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents:
51782
diff
changeset
|
17 |
val ID_bnf: BNF_Def.bnf |
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents:
51782
diff
changeset
|
18 |
val DEADID_bnf: BNF_Def.bnf |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
19 |
|
55706 | 20 |
type comp_cache |
65436 | 21 |
type unfold_set = |
22 |
{map_unfolds: thm list, |
|
23 |
set_unfoldss: thm list list, |
|
71261 | 24 |
rel_unfolds: thm list, |
25 |
pred_unfolds: thm list} |
|
55706 | 26 |
|
27 |
val empty_comp_cache: comp_cache |
|
49502 | 28 |
val empty_unfolds: unfold_set |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
29 |
|
53222 | 30 |
exception BAD_DEAD of typ * typ |
31 |
||
62621 | 32 |
val bnf_of_typ: bool -> BNF_Def.inline_policy -> (binding -> binding) -> |
55703
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
blanchet
parents:
55480
diff
changeset
|
33 |
((string * sort) list list -> (string * sort) list) -> (string * sort) list -> |
55904 | 34 |
(string * sort) list -> typ -> (comp_cache * unfold_set) * local_theory -> |
35 |
(BNF_Def.bnf * (typ list * typ list)) * ((comp_cache * unfold_set) * local_theory) |
|
49014 | 36 |
val default_comp_sort: (string * sort) list list -> (string * sort) list |
59725 | 37 |
val clean_compose_bnf: BNF_Def.inline_policy -> (binding -> binding) -> binding -> BNF_Def.bnf -> |
38 |
BNF_Def.bnf list -> unfold_set * local_theory -> BNF_Def.bnf * (unfold_set * local_theory) |
|
39 |
val kill_bnf: (binding -> binding) -> int -> BNF_Def.bnf -> |
|
40 |
(comp_cache * unfold_set) * local_theory -> |
|
41 |
BNF_Def.bnf * ((comp_cache * unfold_set) * local_theory) |
|
42 |
val lift_bnf: (binding -> binding) -> int -> BNF_Def.bnf -> |
|
43 |
(comp_cache * unfold_set) * local_theory -> |
|
44 |
BNF_Def.bnf * ((comp_cache * unfold_set) * local_theory) |
|
45 |
val permute_bnf: (binding -> binding) -> int list -> int list -> BNF_Def.bnf -> |
|
46 |
(comp_cache * unfold_set) * local_theory -> |
|
47 |
BNF_Def.bnf * ((comp_cache * unfold_set) * local_theory) |
|
48 |
val permute_and_kill_bnf: (binding -> binding) -> int -> int list -> int list -> BNF_Def.bnf -> |
|
49 |
(comp_cache * unfold_set) * local_theory -> |
|
50 |
BNF_Def.bnf * ((comp_cache * unfold_set) * local_theory) |
|
51 |
val lift_and_permute_bnf: (binding -> binding) -> int -> int list -> int list -> BNF_Def.bnf -> |
|
52 |
(comp_cache * unfold_set) * local_theory -> |
|
53 |
BNF_Def.bnf * ((comp_cache * unfold_set) * local_theory) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
54 |
val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list -> |
55904 | 55 |
(''a list list -> ''a list) -> BNF_Def.bnf list -> (comp_cache * unfold_set) * local_theory -> |
56 |
(int list list * ''a list) * (BNF_Def.bnf list * ((comp_cache * unfold_set) * local_theory)) |
|
59725 | 57 |
val compose_bnf: BNF_Def.inline_policy -> (int -> binding -> binding) -> |
58 |
((string * sort) list list -> (string * sort) list) -> BNF_Def.bnf -> BNF_Def.bnf list -> |
|
59 |
typ list -> typ list list -> typ list list -> (comp_cache * unfold_set) * local_theory -> |
|
60 |
(BNF_Def.bnf * (typ list * typ list)) * ((comp_cache * unfold_set) * local_theory) |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
61 |
type absT_info = |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
62 |
{absT: typ, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
63 |
repT: typ, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
64 |
abs: term, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
65 |
rep: term, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
66 |
abs_inject: thm, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
67 |
abs_inverse: thm, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
68 |
type_definition: thm} |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
69 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
70 |
val morph_absT_info: morphism -> absT_info -> absT_info |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
71 |
val mk_absT: theory -> typ -> typ -> typ -> typ |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
72 |
val mk_repT: typ -> typ -> typ -> typ |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
73 |
val mk_abs: typ -> term -> term |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
74 |
val mk_rep: typ -> term -> term |
63837
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
blanchet
parents:
63836
diff
changeset
|
75 |
val seal_bnf: (binding -> binding) -> unfold_set -> binding -> bool -> typ list -> typ list -> |
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
blanchet
parents:
63836
diff
changeset
|
76 |
BNF_Def.bnf -> local_theory -> (BNF_Def.bnf * (typ list * absT_info)) * local_theory |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
77 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
78 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
79 |
structure BNF_Comp : BNF_COMP = |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
80 |
struct |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
81 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
82 |
open BNF_Def |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
83 |
open BNF_Util |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
84 |
open BNF_Tactics |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
85 |
open BNF_Comp_Tactics |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
86 |
|
69593 | 87 |
val typedef_threshold = Attrib.setup_config_int \<^binding>\<open>bnf_typedef_threshold\<close> (K 6); |
59710 | 88 |
|
63800 | 89 |
fun with_typedef_threshold threshold f lthy = |
90 |
lthy |
|
91 |
|> Config.put typedef_threshold threshold |
|
92 |
|> f |
|
93 |
|> Config.put typedef_threshold (Config.get lthy typedef_threshold); |
|
94 |
||
95 |
fun with_typedef_threshold_yield threshold f lthy = |
|
96 |
lthy |
|
97 |
|> Config.put typedef_threshold threshold |
|
98 |
|> f |
|
99 |
||> Config.put typedef_threshold (Config.get lthy typedef_threshold); |
|
100 |
||
69593 | 101 |
val ID_bnf = the (bnf_of \<^context> "BNF_Composition.ID"); |
102 |
val DEADID_bnf = the (bnf_of \<^context> "BNF_Composition.DEADID"); |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
103 |
|
55706 | 104 |
type comp_cache = (bnf * (typ list * typ list)) Typtab.table; |
105 |
||
55904 | 106 |
fun key_of_types s Ts = Type (s, Ts); |
107 |
fun key_of_typess s = key_of_types s o map (key_of_types ""); |
|
108 |
fun typ_of_int n = Type (string_of_int n, []); |
|
109 |
fun typ_of_bnf bnf = |
|
110 |
key_of_typess "" [[T_of_bnf bnf], lives_of_bnf bnf, sort Term_Ord.typ_ord (deads_of_bnf bnf)]; |
|
111 |
||
112 |
fun key_of_kill n bnf = key_of_types "k" [typ_of_int n, typ_of_bnf bnf]; |
|
113 |
fun key_of_lift n bnf = key_of_types "l" [typ_of_int n, typ_of_bnf bnf]; |
|
114 |
fun key_of_permute src dest bnf = |
|
115 |
key_of_types "p" (map typ_of_int src @ map typ_of_int dest @ [typ_of_bnf bnf]); |
|
116 |
fun key_of_compose oDs Dss Ass outer inners = |
|
117 |
key_of_types "c" (map (key_of_typess "") [[oDs], Dss, Ass, [map typ_of_bnf (outer :: inners)]]); |
|
118 |
||
119 |
fun cache_comp_simple key cache (bnf, (unfold_set, lthy)) = |
|
120 |
(bnf, ((Typtab.update (key, (bnf, ([], []))) cache, unfold_set), lthy)); |
|
121 |
||
122 |
fun cache_comp key (bnf_Ds_As, ((cache, unfold_set), lthy)) = |
|
123 |
(bnf_Ds_As, ((Typtab.update (key, bnf_Ds_As) cache, unfold_set), lthy)); |
|
124 |
||
49502 | 125 |
type unfold_set = { |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
126 |
map_unfolds: thm list, |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
127 |
set_unfoldss: thm list list, |
71261 | 128 |
rel_unfolds: thm list, |
129 |
pred_unfolds: thm list |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
130 |
}; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
131 |
|
55706 | 132 |
val empty_comp_cache = Typtab.empty; |
71261 | 133 |
val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = [], pred_unfolds = []}; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
134 |
|
49503 | 135 |
fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new; |
136 |
fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms; |
|
137 |
||
71261 | 138 |
fun add_to_unfolds map sets rel pred |
139 |
{map_unfolds, set_unfoldss, rel_unfolds, pred_unfolds} = |
|
49503 | 140 |
{map_unfolds = add_to_thms map_unfolds map, |
141 |
set_unfoldss = adds_to_thms set_unfoldss sets, |
|
71261 | 142 |
rel_unfolds = add_to_thms rel_unfolds rel, |
143 |
pred_unfolds = add_to_thms pred_unfolds pred}; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
144 |
|
49503 | 145 |
fun add_bnf_to_unfolds bnf = |
71261 | 146 |
add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf) (pred_def_of_bnf bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
147 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
148 |
val bdTN = "bdT"; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
149 |
|
49425 | 150 |
fun mk_killN n = "_kill" ^ string_of_int n; |
151 |
fun mk_liftN n = "_lift" ^ string_of_int n; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
152 |
fun mk_permuteN src dest = |
49425 | 153 |
"_permute_" ^ implode (map string_of_int src) ^ "_" ^ implode (map string_of_int dest); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
154 |
|
55935
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
traytel
parents:
55930
diff
changeset
|
155 |
|
58181 | 156 |
val id_bnf_def = @{thm id_bnf_def}; |
74575 | 157 |
val expand_id_bnf_def = |
158 |
Envir.expand_term_defs dest_Const |
|
159 |
[Thm.prop_of id_bnf_def |> Logic.dest_equals |> apfst dest_Const]; |
|
55937 | 160 |
|
69593 | 161 |
fun is_sum_prod_natLeq (Const (\<^const_name>\<open>csum\<close>, _) $ t $ u) = forall is_sum_prod_natLeq [t, u] |
162 |
| is_sum_prod_natLeq (Const (\<^const_name>\<open>cprod\<close>, _) $ t $ u) = forall is_sum_prod_natLeq [t, u] |
|
163 |
| is_sum_prod_natLeq t = t aconv \<^term>\<open>natLeq\<close>; |
|
55851
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
164 |
|
49502 | 165 |
fun clean_compose_bnf const_policy qualify b outer inners (unfold_set, lthy) = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
166 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
167 |
val olive = live_of_bnf outer; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
168 |
val onwits = nwits_of_bnf outer; |
60207 | 169 |
val odeads = deads_of_bnf outer; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
170 |
val inner = hd inners; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
171 |
val ilive = live_of_bnf inner; |
60207 | 172 |
val ideadss = map deads_of_bnf inners; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
173 |
val inwitss = map nwits_of_bnf inners; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
174 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
175 |
(* TODO: check olive = length inners > 0, |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
176 |
forall inner from inners. ilive = live, |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
177 |
forall inner from inners. idead = dead *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
178 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
179 |
val (oDs, lthy1) = apfst (map TFree) |
60207 | 180 |
(Variable.invent_types (map Type.sort_of_atyp odeads) lthy); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
181 |
val (Dss, lthy2) = apfst (map (map TFree)) |
60207 | 182 |
(fold_map Variable.invent_types (map (map Type.sort_of_atyp) ideadss) lthy1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
183 |
val (Ass, lthy3) = apfst (replicate ilive o map TFree) |
69593 | 184 |
(Variable.invent_types (replicate ilive \<^sort>\<open>type\<close>) lthy2); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
185 |
val As = if ilive > 0 then hd Ass else []; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
186 |
val Ass_repl = replicate olive As; |
55906 | 187 |
val (Bs, names_lthy) = apfst (map TFree) |
69593 | 188 |
(Variable.invent_types (replicate ilive \<^sort>\<open>type\<close>) lthy3); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
189 |
val Bss_repl = replicate olive Bs; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
190 |
|
62324 | 191 |
val (((((fs', Qs'), Ps'), Asets), xs), _) = names_lthy |
52923 | 192 |
|> apfst snd o mk_Frees' "f" (map2 (curry op -->) As Bs) |
49463 | 193 |
||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs) |
62324 | 194 |
||>> apfst snd o mk_Frees' "P" (map mk_pred1T As) |
49456 | 195 |
||>> mk_Frees "A" (map HOLogic.mk_setT As) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
196 |
||>> mk_Frees "x" As; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
197 |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58359
diff
changeset
|
198 |
val CAs = @{map 3} mk_T_of_bnf Dss Ass_repl inners; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
199 |
val CCA = mk_T_of_bnf oDs CAs outer; |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58359
diff
changeset
|
200 |
val CBs = @{map 3} mk_T_of_bnf Dss Bss_repl inners; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
201 |
val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer; |
61760 | 202 |
val inner_setss = |
203 |
@{map 3} mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners; |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58359
diff
changeset
|
204 |
val inner_bds = @{map 3} mk_bd_of_bnf Dss Ass_repl inners; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
205 |
val outer_bd = mk_bd_of_bnf oDs CAs outer; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
206 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
207 |
(*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*) |
49303 | 208 |
val mapx = fold_rev Term.abs fs' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
209 |
(Term.list_comb (mk_map_of_bnf oDs CAs CBs outer, |
49463 | 210 |
map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
211 |
mk_map_of_bnf Ds As Bs) Dss inners)); |
49507 | 212 |
(*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*) |
213 |
val rel = fold_rev Term.abs Qs' |
|
214 |
(Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer, |
|
49463 | 215 |
map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o |
49507 | 216 |
mk_rel_of_bnf Ds As Bs) Dss inners)); |
62324 | 217 |
(*%P1 ... Pn. outer.pred (inner_1.pred P1 ... Pn) ... (inner_m.pred P1 ... Pn)*) |
218 |
val pred = fold_rev Term.abs Ps' |
|
219 |
(Term.list_comb (mk_pred_of_bnf oDs CAs outer, |
|
220 |
map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o |
|
221 |
mk_pred_of_bnf Ds As) Dss inners)); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
222 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
223 |
(*Union o collect {outer.set_1 ... outer.set_m} o outer.map inner_1.set_i ... inner_m.set_i*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
224 |
(*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*) |
49303 | 225 |
fun mk_set i = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
226 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
227 |
val (setTs, T) = `(replicate olive o HOLogic.mk_setT) (nth As i); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
228 |
val outer_set = mk_collect |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
229 |
(mk_sets_of_bnf (replicate olive oDs) (replicate olive setTs) outer) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
230 |
(mk_T_of_bnf oDs setTs outer --> HOLogic.mk_setT T); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
231 |
val inner_sets = map (fn sets => nth sets i) inner_setss; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
232 |
val outer_map = mk_map_of_bnf oDs CAs setTs outer; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
233 |
val map_inner_sets = Term.list_comb (outer_map, inner_sets); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
234 |
val collect_image = mk_collect |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
235 |
(map2 (fn f => fn set => HOLogic.mk_comp (mk_image f, set)) inner_sets outer_sets) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
236 |
(CCA --> HOLogic.mk_setT T); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
237 |
in |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
238 |
(Library.foldl1 HOLogic.mk_comp [mk_Union T, outer_set, map_inner_sets], |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
239 |
HOLogic.mk_comp (mk_Union T, collect_image)) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
240 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
241 |
|
49303 | 242 |
val (sets, sets_alt) = map_split mk_set (0 upto ilive - 1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
243 |
|
55906 | 244 |
fun mk_simplified_set set = |
245 |
let |
|
246 |
val setT = fastype_of set; |
|
69593 | 247 |
val var_set' = Const (\<^const_name>\<open>id_bnf\<close>, setT --> setT) $ Var ((Name.uu, 0), setT); |
55908
e6d570cb0f5e
no 'sorry' so that the schematic variable gets instantiated
blanchet
parents:
55906
diff
changeset
|
248 |
val goal = mk_Trueprop_eq (var_set', set); |
55930
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
traytel
parents:
55908
diff
changeset
|
249 |
fun tac {context = ctxt, prems = _} = |
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
traytel
parents:
55908
diff
changeset
|
250 |
mk_simplified_set_tac ctxt (collect_set_map_of_bnf outer); |
55906 | 251 |
val set'_eq_set = |
57890
1e13f63fb452
use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents:
57632
diff
changeset
|
252 |
Goal.prove (*no sorry*) names_lthy [] [] goal tac |
70494 | 253 |
|> Thm.close_derivation \<^here>; |
55906 | 254 |
val set' = fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of set'_eq_set))); |
255 |
in |
|
256 |
(set', set'_eq_set) |
|
257 |
end; |
|
258 |
||
259 |
val (sets', set'_eq_sets) = |
|
260 |
map_split mk_simplified_set sets |
|
261 |
||> Proof_Context.export names_lthy lthy; |
|
262 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
263 |
(*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*) |
54421 | 264 |
val bd = mk_cprod (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
265 |
|
55851
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
266 |
val (bd', bd_ordIso_natLeq_thm_opt) = |
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
267 |
if is_sum_prod_natLeq bd then |
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
268 |
let |
69593 | 269 |
val bd' = \<^term>\<open>natLeq\<close>; |
55851
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
270 |
val bd_bd' = HOLogic.mk_prod (bd, bd'); |
69593 | 271 |
val ordIso = Const (\<^const_name>\<open>ordIso\<close>, HOLogic.mk_setT (fastype_of bd_bd')); |
57567 | 272 |
val goal = mk_Trueprop_mem (bd_bd', ordIso); |
55851
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
273 |
in |
60757 | 274 |
(bd', SOME (Goal.prove_sorry lthy [] [] goal (bd_ordIso_natLeq_tac o #context) |
70494 | 275 |
|> Thm.close_derivation \<^here>)) |
55851
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
276 |
end |
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
277 |
else |
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
278 |
(bd, NONE); |
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55803
diff
changeset
|
279 |
|
60728 | 280 |
fun map_id0_tac ctxt = |
281 |
mk_comp_map_id0_tac ctxt (map_id0_of_bnf outer) (map_cong0_of_bnf outer) |
|
53270 | 282 |
(map map_id0_of_bnf inners); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
283 |
|
60728 | 284 |
fun map_comp0_tac ctxt = |
285 |
mk_comp_map_comp0_tac ctxt (map_comp0_of_bnf outer) (map_cong0_of_bnf outer) |
|
53287 | 286 |
(map map_comp0_of_bnf inners); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
287 |
|
55906 | 288 |
fun mk_single_set_map0_tac i ctxt = |
289 |
mk_comp_set_map0_tac ctxt (nth set'_eq_sets i) (map_comp0_of_bnf outer) |
|
290 |
(map_cong0_of_bnf outer) (collect_set_map_of_bnf outer) |
|
53289 | 291 |
(map ((fn thms => nth thms i) o set_map0_of_bnf) inners); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
292 |
|
53289 | 293 |
val set_map0_tacs = map mk_single_set_map0_tac (0 upto ilive - 1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
294 |
|
60728 | 295 |
fun bd_card_order_tac ctxt = |
296 |
mk_comp_bd_card_order_tac ctxt (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
297 |
|
60728 | 298 |
fun bd_cinfinite_tac ctxt = |
299 |
mk_comp_bd_cinfinite_tac ctxt (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
300 |
|
75624 | 301 |
fun bd_regularCard_tac ctxt = |
302 |
mk_comp_bd_regularCard_tac ctxt (map bd_regularCard_of_bnf inners) (bd_regularCard_of_bnf outer) |
|
303 |
(map bd_Cinfinite_of_bnf inners) (bd_Cinfinite_of_bnf outer); |
|
304 |
||
49303 | 305 |
val set_alt_thms = |
52059 | 306 |
if Config.get lthy quick_and_dirty then |
49456 | 307 |
[] |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
308 |
else |
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49019
diff
changeset
|
309 |
map (fn goal => |
51551 | 310 |
Goal.prove_sorry lthy [] [] goal |
49714 | 311 |
(fn {context = ctxt, prems = _} => |
51766
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents:
51761
diff
changeset
|
312 |
mk_comp_set_alt_tac ctxt (collect_set_map_of_bnf outer)) |
70494 | 313 |
|> Thm.close_derivation \<^here>) |
63824 | 314 |
(map2 (curry mk_Trueprop_eq) sets sets_alt); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
315 |
|
55906 | 316 |
fun map_cong0_tac ctxt = |
317 |
mk_comp_map_cong0_tac ctxt set'_eq_sets set_alt_thms (map_cong0_of_bnf outer) |
|
318 |
(map map_cong0_of_bnf inners); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
319 |
|
49303 | 320 |
val set_bd_tacs = |
52059 | 321 |
if Config.get lthy quick_and_dirty then |
49669 | 322 |
replicate ilive (K all_tac) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
323 |
else |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
324 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
325 |
val outer_set_bds = set_bd_of_bnf outer; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
326 |
val inner_set_bdss = map set_bd_of_bnf inners; |
75624 | 327 |
val inner_bd_Cinfinites = map bd_Cinfinite_of_bnf inners; |
328 |
val inner_bd_regularCards = map bd_regularCard_of_bnf inners; |
|
49303 | 329 |
fun single_set_bd_thm i j = |
75624 | 330 |
@{thm comp_single_set_bd_strict} OF [nth inner_bd_Cinfinites j, nth inner_bd_regularCards j, |
331 |
bd_Cinfinite_of_bnf outer, bd_regularCard_of_bnf outer, nth (nth inner_set_bdss j) i, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
332 |
nth outer_set_bds j] |
75624 | 333 |
fun single_bd_Cinfinite i = @{thm Cinfinite_cprod2} OF [ |
334 |
@{thm Cinfinite_Cnotzero} OF [bd_Cinfinite_of_bnf outer], |
|
335 |
nth inner_bd_Cinfinites i |
|
336 |
] |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
337 |
val single_set_bd_thmss = |
49303 | 338 |
map ((fn f => map f (0 upto olive - 1)) o single_set_bd_thm) (0 upto ilive - 1); |
75624 | 339 |
val Gbd_Fbd_Cinfinites = map single_bd_Cinfinite (0 upto length inners - 1) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
340 |
in |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58359
diff
changeset
|
341 |
@{map 3} (fn set'_eq_set => fn set_alt => fn single_set_bds => fn ctxt => |
75624 | 342 |
mk_comp_set_bd_tac ctxt set'_eq_set bd_ordIso_natLeq_thm_opt set_alt single_set_bds |
343 |
Gbd_Fbd_Cinfinites) |
|
55906 | 344 |
set'_eq_sets set_alt_thms single_set_bd_thmss |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
345 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
346 |
|
49303 | 347 |
val in_alt_thm = |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
348 |
let |
49303 | 349 |
val inx = mk_in Asets sets CCA; |
350 |
val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA; |
|
351 |
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); |
|
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
352 |
in |
51551 | 353 |
Goal.prove_sorry lthy [] [] goal |
49714 | 354 |
(fn {context = ctxt, prems = _} => mk_comp_in_alt_tac ctxt set_alt_thms) |
70494 | 355 |
|> Thm.close_derivation \<^here> |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
356 |
end; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
357 |
|
60728 | 358 |
fun le_rel_OO_tac ctxt = mk_le_rel_OO_tac ctxt (le_rel_OO_of_bnf outer) (rel_mono_of_bnf outer) |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54742
diff
changeset
|
359 |
(map le_rel_OO_of_bnf inners); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
360 |
|
55906 | 361 |
fun rel_OO_Grp_tac ctxt = |
49456 | 362 |
let |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
363 |
val outer_rel_Grp = rel_Grp_of_bnf outer RS sym; |
49463 | 364 |
val thm = |
62324 | 365 |
trans OF [in_alt_thm RS @{thm OO_Grp_cong}, |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
366 |
trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
367 |
[trans OF [outer_rel_Grp RS @{thm arg_cong[of _ _ conversep]}, |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
368 |
rel_conversep_of_bnf outer RS sym], outer_rel_Grp], |
61242 | 369 |
trans OF [rel_OO_of_bnf outer RS sym, rel_cong0_of_bnf outer OF |
62324 | 370 |
(map (fn bnf => rel_OO_Grp_of_bnf bnf RS sym) inners)]]] RS sym; |
49456 | 371 |
in |
60728 | 372 |
unfold_thms_tac ctxt set'_eq_sets THEN rtac ctxt thm 1 |
49463 | 373 |
end; |
49456 | 374 |
|
62324 | 375 |
fun pred_set_tac ctxt = |
376 |
let |
|
377 |
val pred_alt = unfold_thms ctxt @{thms Ball_Collect} |
|
378 |
(trans OF [pred_cong0_of_bnf outer OF map pred_set_of_bnf inners, pred_set_of_bnf outer]); |
|
379 |
val in_alt = in_alt_thm RS @{thm Collect_inj} RS sym; |
|
380 |
in |
|
381 |
unfold_thms_tac ctxt (@{thm Ball_Collect} :: set'_eq_sets) THEN |
|
382 |
HEADGOAL (rtac ctxt (trans OF [pred_alt, in_alt])) |
|
383 |
end |
|
384 |
||
53289 | 385 |
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac |
75624 | 386 |
bd_cinfinite_tac bd_regularCard_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac pred_set_tac; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
387 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
388 |
val outer_wits = mk_wits_of_bnf (replicate onwits oDs) (replicate onwits CAs) outer; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
389 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
390 |
val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I))) |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58359
diff
changeset
|
391 |
(@{map 3} (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As)) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
392 |
Dss inwitss inners); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
393 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
394 |
val inner_witsss = map (map (nth inner_witss) o fst) outer_wits; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
395 |
|
49303 | 396 |
val wits = (inner_witsss, (map (single o snd) outer_wits)) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
397 |
|-> map2 (fold (map_product (fn iwit => fn owit => owit $ iwit))) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
398 |
|> flat |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
399 |
|> map (`(fn t => Term.add_frees t [])) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
400 |
|> minimize_wits |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
401 |
|> map (fn (frees, t) => fold absfree frees t); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
402 |
|
55197 | 403 |
fun wit_tac ctxt = |
55906 | 404 |
mk_comp_wit_tac ctxt set'_eq_sets (wit_thms_of_bnf outer) (collect_set_map_of_bnf outer) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
405 |
(maps wit_thms_of_bnf inners); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
406 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
407 |
val (bnf', lthy') = |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
408 |
bnf_def const_policy (K Dont_Note) true qualify tacs wit_tac (SOME (oDs @ flat Dss)) |
62324 | 409 |
Binding.empty Binding.empty Binding.empty [] |
410 |
(((((((b, CCA), mapx), sets'), bd'), wits), SOME rel), SOME pred) lthy; |
|
55906 | 411 |
|
412 |
val phi = |
|
58181 | 413 |
Morphism.thm_morphism "BNF" (unfold_thms lthy' [id_bnf_def]) |
414 |
$> Morphism.term_morphism "BNF" expand_id_bnf_def; |
|
55906 | 415 |
|
416 |
val bnf'' = morph_bnf phi bnf'; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
417 |
in |
55906 | 418 |
(bnf'', (add_bnf_to_unfolds bnf'' unfold_set, lthy')) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
419 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
420 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
421 |
(* Killing live variables *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
422 |
|
55904 | 423 |
fun raw_kill_bnf qualify n bnf (accum as (unfold_set, lthy)) = |
424 |
if n = 0 then (bnf, accum) else |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
425 |
let |
49425 | 426 |
val b = Binding.suffix_name (mk_killN n) (name_of_bnf bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
427 |
val live = live_of_bnf bnf; |
60207 | 428 |
val deads = deads_of_bnf bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
429 |
val nwits = nwits_of_bnf bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
430 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
431 |
(* TODO: check 0 < n <= live *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
432 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
433 |
val (Ds, lthy1) = apfst (map TFree) |
60207 | 434 |
(Variable.invent_types (map Type.sort_of_atyp deads) lthy); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
435 |
val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree) |
69593 | 436 |
(Variable.invent_types (replicate live \<^sort>\<open>type\<close>) lthy1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
437 |
val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree) |
69593 | 438 |
(Variable.invent_types (replicate (live - n) \<^sort>\<open>type\<close>) lthy2); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
439 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
440 |
val ((Asets, lives), _(*names_lthy*)) = lthy |
49456 | 441 |
|> mk_Frees "A" (map HOLogic.mk_setT (drop n As)) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
442 |
||>> mk_Frees "x" (drop n As); |
69593 | 443 |
val xs = map (fn T => Const (\<^const_name>\<open>undefined\<close>, T)) killedAs @ lives; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
444 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
445 |
val T = mk_T_of_bnf Ds As bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
446 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
447 |
(*bnf.map id ... id*) |
49303 | 448 |
val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs); |
49507 | 449 |
(*bnf.rel (op =) ... (op =)*) |
450 |
val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs); |
|
62324 | 451 |
(*bnf.pred (%_. True) ... (%_ True)*) |
452 |
val pred = Term.list_comb (mk_pred_of_bnf Ds As bnf, |
|
69593 | 453 |
map (fn T => Term.absdummy T \<^term>\<open>True\<close>) killedAs); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
454 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
455 |
val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf; |
49303 | 456 |
val sets = drop n bnf_sets; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
457 |
|
55707
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
traytel
parents:
55706
diff
changeset
|
458 |
val bd = mk_bd_of_bnf Ds As bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
459 |
|
60728 | 460 |
fun map_id0_tac ctxt = rtac ctxt (map_id0_of_bnf bnf) 1; |
55197 | 461 |
fun map_comp0_tac ctxt = |
55067 | 462 |
unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: |
60728 | 463 |
@{thms comp_assoc id_comp comp_id}) THEN rtac ctxt refl 1; |
55197 | 464 |
fun map_cong0_tac ctxt = |
51761
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents:
51758
diff
changeset
|
465 |
mk_kill_map_cong0_tac ctxt n (live - n) (map_cong0_of_bnf bnf); |
60728 | 466 |
val set_map0_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) (drop n (set_map0_of_bnf bnf)); |
467 |
fun bd_card_order_tac ctxt = rtac ctxt (bd_card_order_of_bnf bnf) 1; |
|
468 |
fun bd_cinfinite_tac ctxt = rtac ctxt (bd_cinfinite_of_bnf bnf) 1; |
|
75624 | 469 |
fun bd_regularCard_tac ctxt = rtac ctxt (bd_regularCard_of_bnf bnf) 1; |
60728 | 470 |
val set_bd_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) (drop n (set_bd_of_bnf bnf)); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
471 |
|
49303 | 472 |
val in_alt_thm = |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
473 |
let |
49303 | 474 |
val inx = mk_in Asets sets T; |
475 |
val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T; |
|
476 |
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); |
|
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
477 |
in |
60728 | 478 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
70494 | 479 |
kill_in_alt_tac ctxt) |> Thm.close_derivation \<^here> |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
480 |
end; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
481 |
|
55197 | 482 |
fun le_rel_OO_tac ctxt = |
60728 | 483 |
EVERY' [rtac ctxt @{thm ord_le_eq_trans}, rtac ctxt (le_rel_OO_of_bnf bnf)] 1 THEN |
484 |
unfold_thms_tac ctxt @{thms eq_OO} THEN rtac ctxt refl 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
485 |
|
60728 | 486 |
fun rel_OO_Grp_tac ctxt = |
49456 | 487 |
let |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
488 |
val rel_Grp = rel_Grp_of_bnf bnf RS sym |
49463 | 489 |
val thm = |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
490 |
(trans OF [in_alt_thm RS @{thm OO_Grp_cong}, |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
491 |
trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
492 |
[trans OF [rel_Grp RS @{thm arg_cong[of _ _ conversep]}, |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
493 |
rel_conversep_of_bnf bnf RS sym], rel_Grp], |
61242 | 494 |
trans OF [rel_OO_of_bnf bnf RS sym, rel_cong0_of_bnf bnf OF |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51837
diff
changeset
|
495 |
(replicate n @{thm trans[OF Grp_UNIV_id[OF refl] eq_alt[symmetric]]} @ |
52660 | 496 |
replicate (live - n) @{thm Grp_fst_snd})]]] RS sym); |
49456 | 497 |
in |
60728 | 498 |
rtac ctxt thm 1 |
49456 | 499 |
end; |
500 |
||
62324 | 501 |
fun pred_set_tac ctxt = mk_simple_pred_set_tac ctxt (pred_set_of_bnf bnf) in_alt_thm; |
502 |
||
53289 | 503 |
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac |
75624 | 504 |
bd_cinfinite_tac bd_regularCard_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac pred_set_tac; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
505 |
|
49303 | 506 |
val bnf_wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
507 |
|
49303 | 508 |
val wits = map (fn t => fold absfree (Term.add_frees t []) t) |
509 |
(map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
510 |
|
60752 | 511 |
fun wit_tac ctxt = mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
512 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
513 |
val (bnf', lthy') = |
58297 | 514 |
bnf_def Smart_Inline (K Dont_Note) true qualify tacs wit_tac (SOME (Ds @ killedAs)) |
62324 | 515 |
Binding.empty Binding.empty Binding.empty [] |
516 |
(((((((b, T), mapx), sets), bd), wits), SOME rel), SOME pred) lthy; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
517 |
in |
49503 | 518 |
(bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy')) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
519 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
520 |
|
55904 | 521 |
fun kill_bnf qualify n bnf (accum as ((cache, unfold_set), lthy)) = |
522 |
let val key = key_of_kill n bnf in |
|
523 |
(case Typtab.lookup cache key of |
|
524 |
SOME (bnf, _) => (bnf, accum) |
|
525 |
| NONE => cache_comp_simple key cache (raw_kill_bnf qualify n bnf (unfold_set, lthy))) |
|
526 |
end; |
|
527 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
528 |
(* Adding dummy live variables *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
529 |
|
55904 | 530 |
fun raw_lift_bnf qualify n bnf (accum as (unfold_set, lthy)) = |
531 |
if n = 0 then (bnf, accum) else |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
532 |
let |
49425 | 533 |
val b = Binding.suffix_name (mk_liftN n) (name_of_bnf bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
534 |
val live = live_of_bnf bnf; |
60207 | 535 |
val deads = deads_of_bnf bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
536 |
val nwits = nwits_of_bnf bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
537 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
538 |
(* TODO: check 0 < n *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
539 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
540 |
val (Ds, lthy1) = apfst (map TFree) |
60207 | 541 |
(Variable.invent_types (map Type.sort_of_atyp deads) lthy); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
542 |
val ((newAs, As), lthy2) = apfst (chop n o map TFree) |
69593 | 543 |
(Variable.invent_types (replicate (n + live) \<^sort>\<open>type\<close>) lthy1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
544 |
val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree) |
69593 | 545 |
(Variable.invent_types (replicate (n + live) \<^sort>\<open>type\<close>) lthy2); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
546 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
547 |
val (Asets, _(*names_lthy*)) = lthy |
49456 | 548 |
|> mk_Frees "A" (map HOLogic.mk_setT (newAs @ As)); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
549 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
550 |
val T = mk_T_of_bnf Ds As bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
551 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
552 |
(*%f1 ... fn. bnf.map*) |
49303 | 553 |
val mapx = |
52923 | 554 |
fold_rev Term.absdummy (map2 (curry op -->) newAs newBs) (mk_map_of_bnf Ds As Bs bnf); |
49507 | 555 |
(*%Q1 ... Qn. bnf.rel*) |
556 |
val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf); |
|
62324 | 557 |
(*%P1 ... Pn. bnf.pred*) |
558 |
val pred = fold_rev Term.absdummy (map mk_pred1T newAs) (mk_pred_of_bnf Ds As bnf); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
559 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
560 |
val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf; |
49303 | 561 |
val sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
562 |
|
49303 | 563 |
val bd = mk_bd_of_bnf Ds As bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
564 |
|
60728 | 565 |
fun map_id0_tac ctxt = rtac ctxt (map_id0_of_bnf bnf) 1; |
55197 | 566 |
fun map_comp0_tac ctxt = |
55067 | 567 |
unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: |
60728 | 568 |
@{thms comp_assoc id_comp comp_id}) THEN rtac ctxt refl 1; |
55197 | 569 |
fun map_cong0_tac ctxt = |
60728 | 570 |
rtac ctxt (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); |
53289 | 571 |
val set_map0_tacs = |
52059 | 572 |
if Config.get lthy quick_and_dirty then |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
573 |
replicate (n + live) (K all_tac) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
574 |
else |
60728 | 575 |
replicate n empty_natural_tac @ |
576 |
map (fn thm => fn ctxt => rtac ctxt thm 1) (set_map0_of_bnf bnf); |
|
577 |
fun bd_card_order_tac ctxt = rtac ctxt (bd_card_order_of_bnf bnf) 1; |
|
578 |
fun bd_cinfinite_tac ctxt = rtac ctxt (bd_cinfinite_of_bnf bnf) 1; |
|
75624 | 579 |
fun bd_regularCard_tac ctxt = rtac ctxt (bd_regularCard_of_bnf bnf) 1; |
49303 | 580 |
val set_bd_tacs = |
52059 | 581 |
if Config.get lthy quick_and_dirty then |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
582 |
replicate (n + live) (K all_tac) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
583 |
else |
75624 | 584 |
replicate n (fn ctxt => mk_lift_set_bd_tac ctxt (bd_Cinfinite_of_bnf bnf)) @ |
60728 | 585 |
(map (fn thm => fn ctxt => rtac ctxt thm 1) (set_bd_of_bnf bnf)); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
586 |
|
49303 | 587 |
val in_alt_thm = |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
588 |
let |
49303 | 589 |
val inx = mk_in Asets sets T; |
590 |
val in_alt = mk_in (drop n Asets) bnf_sets T; |
|
591 |
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); |
|
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
592 |
in |
60728 | 593 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => lift_in_alt_tac ctxt) |
70494 | 594 |
|> Thm.close_derivation \<^here> |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
595 |
end; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
596 |
|
60728 | 597 |
fun le_rel_OO_tac ctxt = rtac ctxt (le_rel_OO_of_bnf bnf) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
598 |
|
60728 | 599 |
fun rel_OO_Grp_tac ctxt = mk_simple_rel_OO_Grp_tac ctxt (rel_OO_Grp_of_bnf bnf) in_alt_thm; |
49456 | 600 |
|
62324 | 601 |
fun pred_set_tac ctxt = mk_simple_pred_set_tac ctxt (pred_set_of_bnf bnf) in_alt_thm; |
602 |
||
53289 | 603 |
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac |
75624 | 604 |
bd_cinfinite_tac bd_regularCard_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac pred_set_tac; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
605 |
|
49303 | 606 |
val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
607 |
|
60752 | 608 |
fun wit_tac ctxt = mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
609 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
610 |
val (bnf', lthy') = |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
611 |
bnf_def Smart_Inline (K Dont_Note) true qualify tacs wit_tac (SOME Ds) Binding.empty |
62324 | 612 |
Binding.empty Binding.empty [] |
613 |
(((((((b, T), mapx), sets), bd), wits), SOME rel), SOME pred) lthy; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
614 |
in |
49503 | 615 |
(bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy')) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
616 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
617 |
|
55904 | 618 |
fun lift_bnf qualify n bnf (accum as ((cache, unfold_set), lthy)) = |
619 |
let val key = key_of_lift n bnf in |
|
620 |
(case Typtab.lookup cache key of |
|
621 |
SOME (bnf, _) => (bnf, accum) |
|
622 |
| NONE => cache_comp_simple key cache (raw_lift_bnf qualify n bnf (unfold_set, lthy))) |
|
623 |
end; |
|
624 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
625 |
(* Changing the order of live variables *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
626 |
|
55904 | 627 |
fun raw_permute_bnf qualify src dest bnf (accum as (unfold_set, lthy)) = |
628 |
if src = dest then (bnf, accum) else |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
629 |
let |
68960
b85d509e7cbf
smash position to avoid position of other file "~~/src/HOL/BNF_Composition.thy" (due to "bnf ID"), e.g. relevant for "HOL-Nominal-Examples.Class1";
wenzelm
parents:
67091
diff
changeset
|
630 |
val b = Binding.suffix_name (mk_permuteN src dest) (name_of_bnf bnf) |> Binding.reset_pos; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
631 |
val live = live_of_bnf bnf; |
60207 | 632 |
val deads = deads_of_bnf bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
633 |
val nwits = nwits_of_bnf bnf; |
55480
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
blanchet
parents:
55197
diff
changeset
|
634 |
|
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
blanchet
parents:
55197
diff
changeset
|
635 |
fun permute xs = permute_like_unique (op =) src dest xs; |
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
blanchet
parents:
55197
diff
changeset
|
636 |
fun unpermute xs = permute_like_unique (op =) dest src xs; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
637 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
638 |
val (Ds, lthy1) = apfst (map TFree) |
60207 | 639 |
(Variable.invent_types (map Type.sort_of_atyp deads) lthy); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
640 |
val (As, lthy2) = apfst (map TFree) |
69593 | 641 |
(Variable.invent_types (replicate live \<^sort>\<open>type\<close>) lthy1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
642 |
val (Bs, _(*lthy3*)) = apfst (map TFree) |
69593 | 643 |
(Variable.invent_types (replicate live \<^sort>\<open>type\<close>) lthy2); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
644 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
645 |
val (Asets, _(*names_lthy*)) = lthy |
49456 | 646 |
|> mk_Frees "A" (map HOLogic.mk_setT (permute As)); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
647 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
648 |
val T = mk_T_of_bnf Ds As bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
649 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
650 |
(*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*) |
49303 | 651 |
val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs)) |
53038 | 652 |
(Term.list_comb (mk_map_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0)))); |
49507 | 653 |
(*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*) |
654 |
val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs)) |
|
53038 | 655 |
(Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0)))); |
62324 | 656 |
(*%P(1) ... P(n). bnf.pred P\<sigma>(1) ... P\<sigma>(n)*) |
657 |
val pred = fold_rev Term.absdummy (permute (map mk_pred1T As)) |
|
658 |
(Term.list_comb (mk_pred_of_bnf Ds As bnf, unpermute (map Bound (live - 1 downto 0)))); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
659 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
660 |
val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf; |
49303 | 661 |
val sets = permute bnf_sets; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
662 |
|
49303 | 663 |
val bd = mk_bd_of_bnf Ds As bnf; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
664 |
|
60728 | 665 |
fun map_id0_tac ctxt = rtac ctxt (map_id0_of_bnf bnf) 1; |
666 |
fun map_comp0_tac ctxt = rtac ctxt (map_comp0_of_bnf bnf) 1; |
|
55197 | 667 |
fun map_cong0_tac ctxt = |
60728 | 668 |
rtac ctxt (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); |
669 |
val set_map0_tacs = permute (map (fn thm => fn ctxt => rtac ctxt thm 1) (set_map0_of_bnf bnf)); |
|
670 |
fun bd_card_order_tac ctxt = rtac ctxt (bd_card_order_of_bnf bnf) 1; |
|
671 |
fun bd_cinfinite_tac ctxt = rtac ctxt (bd_cinfinite_of_bnf bnf) 1; |
|
75624 | 672 |
fun bd_regularCard_tac ctxt = rtac ctxt (bd_regularCard_of_bnf bnf) 1; |
60728 | 673 |
val set_bd_tacs = permute (map (fn thm => fn ctxt => rtac ctxt thm 1) (set_bd_of_bnf bnf)); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
674 |
|
49303 | 675 |
val in_alt_thm = |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
676 |
let |
49303 | 677 |
val inx = mk_in Asets sets T; |
53038 | 678 |
val in_alt = mk_in (unpermute Asets) bnf_sets T; |
49303 | 679 |
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
680 |
in |
60728 | 681 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
682 |
mk_permute_in_alt_tac ctxt src dest) |
|
70494 | 683 |
|> Thm.close_derivation \<^here> |
49155
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents:
49123
diff
changeset
|
684 |
end; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
685 |
|
60728 | 686 |
fun le_rel_OO_tac ctxt = rtac ctxt (le_rel_OO_of_bnf bnf) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
687 |
|
60728 | 688 |
fun rel_OO_Grp_tac ctxt = mk_simple_rel_OO_Grp_tac ctxt (rel_OO_Grp_of_bnf bnf) in_alt_thm; |
49456 | 689 |
|
62324 | 690 |
fun pred_set_tac ctxt = mk_simple_pred_set_tac ctxt (pred_set_of_bnf bnf) in_alt_thm; |
691 |
||
53289 | 692 |
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac |
75624 | 693 |
bd_cinfinite_tac bd_regularCard_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac pred_set_tac; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
694 |
|
49303 | 695 |
val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
696 |
|
60752 | 697 |
fun wit_tac ctxt = mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
698 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
699 |
val (bnf', lthy') = |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
700 |
bnf_def Smart_Inline (K Dont_Note) true qualify tacs wit_tac (SOME Ds) Binding.empty |
62324 | 701 |
Binding.empty Binding.empty [] |
702 |
(((((((b, T), mapx), sets), bd), wits), SOME rel), SOME pred) lthy; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
703 |
in |
49503 | 704 |
(bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy')) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
705 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
706 |
|
55904 | 707 |
fun permute_bnf qualify src dest bnf (accum as ((cache, unfold_set), lthy)) = |
708 |
let val key = key_of_permute src dest bnf in |
|
709 |
(case Typtab.lookup cache key of |
|
710 |
SOME (bnf, _) => (bnf, accum) |
|
711 |
| NONE => cache_comp_simple key cache (raw_permute_bnf qualify src dest bnf (unfold_set, lthy))) |
|
712 |
end; |
|
713 |
||
49014 | 714 |
(* Composition pipeline *) |
715 |
||
59725 | 716 |
fun permute_and_kill_bnf qualify n src dest bnf = |
55703
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
blanchet
parents:
55480
diff
changeset
|
717 |
permute_bnf qualify src dest bnf |
49304 | 718 |
#> uncurry (kill_bnf qualify n); |
49014 | 719 |
|
59725 | 720 |
fun lift_and_permute_bnf qualify n src dest bnf = |
55703
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
blanchet
parents:
55480
diff
changeset
|
721 |
lift_bnf qualify n bnf |
49014 | 722 |
#> uncurry (permute_bnf qualify src dest); |
723 |
||
58332 | 724 |
fun normalize_bnfs qualify Ass Ds flatten_tyargs bnfs accum = |
49014 | 725 |
let |
726 |
val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass; |
|
52985 | 727 |
val kill_poss = map (find_indices op = Ds) Ass; |
728 |
val live_poss = map2 (subtract op =) kill_poss before_kill_src; |
|
49014 | 729 |
val before_kill_dest = map2 append kill_poss live_poss; |
730 |
val kill_ns = map length kill_poss; |
|
55904 | 731 |
val (inners', accum') = |
59725 | 732 |
@{fold_map 5} (permute_and_kill_bnf o qualify) |
733 |
(if length bnfs = 1 then [0] else 1 upto length bnfs) |
|
55904 | 734 |
kill_ns before_kill_src before_kill_dest bnfs accum; |
49014 | 735 |
|
736 |
val Ass' = map2 (map o nth) Ass live_poss; |
|
58332 | 737 |
val As = flatten_tyargs Ass'; |
49014 | 738 |
val after_lift_dest = replicate (length Ass') (0 upto (length As - 1)); |
739 |
val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass'; |
|
52985 | 740 |
val new_poss = map2 (subtract op =) old_poss after_lift_dest; |
49014 | 741 |
val after_lift_src = map2 append new_poss old_poss; |
742 |
val lift_ns = map (fn xs => length As - length xs) Ass'; |
|
743 |
in |
|
59725 | 744 |
((kill_poss, As), @{fold_map 5} (lift_and_permute_bnf o qualify) |
55703
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
blanchet
parents:
55480
diff
changeset
|
745 |
(if length bnfs = 1 then [0] else 1 upto length bnfs) |
55904 | 746 |
lift_ns after_lift_src after_lift_dest inners' accum') |
49014 | 747 |
end; |
748 |
||
749 |
fun default_comp_sort Ass = |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58959
diff
changeset
|
750 |
Library.sort (Term_Ord.typ_ord o apply2 TFree) (fold (fold (insert (op =))) Ass []); |
49014 | 751 |
|
58332 | 752 |
fun raw_compose_bnf const_policy qualify flatten_tyargs outer inners oDs Dss tfreess accum = |
49014 | 753 |
let |
49425 | 754 |
val b = name_of_bnf outer; |
49014 | 755 |
|
49121 | 756 |
val Ass = map (map Term.dest_TFree) tfreess; |
49014 | 757 |
val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) []; |
758 |
||
55904 | 759 |
val ((kill_poss, As), (inners', ((cache', unfold_set'), lthy'))) = |
58332 | 760 |
normalize_bnfs qualify Ass Ds flatten_tyargs inners accum; |
49014 | 761 |
|
61760 | 762 |
val Ds = |
763 |
oDs @ flat (@{map 3} (uncurry append oo curry swap oo map o nth) tfreess kill_poss Dss); |
|
49014 | 764 |
val As = map TFree As; |
765 |
in |
|
49425 | 766 |
apfst (rpair (Ds, As)) |
55904 | 767 |
(apsnd (apfst (pair cache')) |
768 |
(clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy'))) |
|
769 |
end; |
|
770 |
||
58332 | 771 |
fun compose_bnf const_policy qualify flatten_tyargs outer inners oDs Dss tfreess |
772 |
(accum as ((cache, _), _)) = |
|
55904 | 773 |
let val key = key_of_compose oDs Dss tfreess outer inners in |
774 |
(case Typtab.lookup cache key of |
|
775 |
SOME bnf_Ds_As => (bnf_Ds_As, accum) |
|
776 |
| NONE => |
|
58332 | 777 |
cache_comp key |
778 |
(raw_compose_bnf const_policy qualify flatten_tyargs outer inners oDs Dss tfreess accum)) |
|
49014 | 779 |
end; |
780 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
781 |
(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
782 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
783 |
type absT_info = |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
784 |
{absT: typ, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
785 |
repT: typ, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
786 |
abs: term, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
787 |
rep: term, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
788 |
abs_inject: thm, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
789 |
abs_inverse: thm, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
790 |
type_definition: thm}; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
791 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
792 |
fun morph_absT_info phi |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
793 |
{absT, repT, abs, rep, abs_inject, abs_inverse, type_definition} = |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
794 |
{absT = Morphism.typ phi absT, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
795 |
repT = Morphism.typ phi repT, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
796 |
abs = Morphism.term phi abs, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
797 |
rep = Morphism.term phi rep, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
798 |
abs_inject = Morphism.thm phi abs_inject, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
799 |
abs_inverse = Morphism.thm phi abs_inverse, |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
800 |
type_definition = Morphism.thm phi type_definition}; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
801 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
802 |
fun mk_absT thy repT absT repU = |
55900
21aa30ea6806
propagate the exception that is expected later on
traytel
parents:
55856
diff
changeset
|
803 |
let |
56634 | 804 |
val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (repT, repU) Vartab.empty) []; |
55900
21aa30ea6806
propagate the exception that is expected later on
traytel
parents:
55856
diff
changeset
|
805 |
in Term.typ_subst_TVars rho absT end |
21aa30ea6806
propagate the exception that is expected later on
traytel
parents:
55856
diff
changeset
|
806 |
handle Type.TYPE_MATCH => raise Term.TYPE ("mk_absT", [repT, absT, repU], []); |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
807 |
|
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
808 |
fun mk_repT absT repT absU = |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
809 |
if absT = repT then absU |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
810 |
else |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
811 |
(case (absT, absU) of |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
812 |
(Type (C, Ts), Type (C', Us)) => |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
813 |
if C = C' then Term.typ_subst_atomic (Ts ~~ Us) repT |
62684
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
traytel
parents:
62621
diff
changeset
|
814 |
else raise Term.TYPE ("mk_repT", [absT, repT, absU], []) |
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
traytel
parents:
62621
diff
changeset
|
815 |
| _ => raise Term.TYPE ("mk_repT", [absT, repT, absU], [])); |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
816 |
|
69593 | 817 |
fun mk_abs_or_rep _ absU (Const (\<^const_name>\<open>id_bnf\<close>, _)) = |
818 |
Const (\<^const_name>\<open>id_bnf\<close>, absU --> absU) |
|
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
819 |
| mk_abs_or_rep getT (Type (_, Us)) abs = |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
75624
diff
changeset
|
820 |
let val Ts = dest_Type_args (getT (fastype_of abs)) |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
821 |
in Term.subst_atomic_types (Ts ~~ Us) abs end; |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
822 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
823 |
val mk_abs = mk_abs_or_rep range_type; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
824 |
val mk_rep = mk_abs_or_rep domain_type; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
825 |
|
63802 | 826 |
fun maybe_typedef force_out_of_line (b, As, mx) set opt_morphs tac lthy = |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
827 |
let |
63802 | 828 |
val threshold = Config.get lthy typedef_threshold; |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
829 |
val repT = HOLogic.dest_setT (fastype_of set); |
59710 | 830 |
val out_of_line = force_out_of_line orelse |
59994 | 831 |
(threshold >= 0 andalso Term.size_of_typ repT >= threshold); |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
832 |
in |
58332 | 833 |
if out_of_line then |
63802 | 834 |
typedef (b, As, mx) set opt_morphs tac lthy |
835 |
|>> (fn (T_name, ({Rep_name, Abs_name, ...}, |
|
58332 | 836 |
{type_definition, Abs_inverse, Abs_inject, Abs_cases, ...}) : Typedef.info) => |
837 |
(Type (T_name, map TFree As), |
|
838 |
(Rep_name, Abs_name, type_definition, Abs_inverse, Abs_inject, Abs_cases))) |
|
839 |
else |
|
63802 | 840 |
((repT, |
69593 | 841 |
(\<^const_name>\<open>id_bnf\<close>, \<^const_name>\<open>id_bnf\<close>, |
58181 | 842 |
@{thm type_definition_id_bnf_UNIV}, |
843 |
@{thm type_definition.Abs_inverse[OF type_definition_id_bnf_UNIV]}, |
|
844 |
@{thm type_definition.Abs_inject[OF type_definition_id_bnf_UNIV]}, |
|
63802 | 845 |
@{thm type_definition.Abs_cases[OF type_definition_id_bnf_UNIV]})), lthy) |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
846 |
end; |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
847 |
|
63837
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
blanchet
parents:
63836
diff
changeset
|
848 |
fun seal_bnf qualify (unfold_set : unfold_set) b force_out_of_line Ds all_Ds bnf lthy = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
849 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
850 |
val live = live_of_bnf bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
851 |
val nwits = nwits_of_bnf bnf; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
852 |
|
58332 | 853 |
val ((As, As'), lthy1) = apfst (`(map TFree)) |
69593 | 854 |
(Variable.invent_types (replicate live \<^sort>\<open>type\<close>) (fold Variable.declare_typ all_Ds lthy)); |
855 |
val (Bs, _) = apfst (map TFree) (Variable.invent_types (replicate live \<^sort>\<open>type\<close>) lthy1); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
856 |
|
62324 | 857 |
val ((((fs, fs'), (Rs, Rs')), (Ps, Ps')), _(*names_lthy*)) = lthy |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
858 |
|> mk_Frees' "f" (map2 (curry op -->) As Bs) |
62324 | 859 |
||>> mk_Frees' "R" (map2 mk_pred2T As Bs) |
860 |
||>> mk_Frees' "P" (map mk_pred1T As); |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
861 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
862 |
val repTA = mk_T_of_bnf Ds As bnf; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
863 |
val T_bind = qualify b; |
63836 | 864 |
val repTA_tfrees = Term.add_tfreesT repTA []; |
63837
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
blanchet
parents:
63836
diff
changeset
|
865 |
val all_TA_params_in_order = fold_rev Term.add_tfreesT all_Ds As'; |
59821
fd3a7692e083
preserve order of type arguments in pre-FP BNF typedef
blanchet
parents:
59820
diff
changeset
|
866 |
val TA_params = |
63802 | 867 |
(if force_out_of_line then all_TA_params_in_order |
63836 | 868 |
else inter (op =) repTA_tfrees all_TA_params_in_order); |
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
869 |
val ((TA, (Rep_name, Abs_name, type_definition, Abs_inverse, Abs_inject, _)), lthy) = |
63802 | 870 |
maybe_typedef force_out_of_line (T_bind, TA_params, NoSyn) (HOLogic.mk_UNIV repTA) NONE |
60728 | 871 |
(fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
872 |
|
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
873 |
val repTB = mk_T_of_bnf Ds Bs bnf; |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
874 |
val TB = Term.typ_subst_atomic (As ~~ Bs) TA; |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
875 |
val RepA = Const (Rep_name, TA --> repTA); |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
876 |
val RepB = Const (Rep_name, TB --> repTB); |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
877 |
val AbsA = Const (Abs_name, repTA --> TA); |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
878 |
val AbsB = Const (Abs_name, repTB --> TB); |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
879 |
val Abs_inject' = Abs_inject OF @{thms UNIV_I UNIV_I}; |
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
880 |
val Abs_inverse' = Abs_inverse OF @{thms UNIV_I}; |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
881 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
882 |
val absT_info = {absT = TA, repT = repTA, abs = AbsA, rep = RepA, abs_inject = Abs_inject', |
55854
ee270328a781
make 'typedef' optional, depending on size of original type
blanchet
parents:
55853
diff
changeset
|
883 |
abs_inverse = Abs_inverse', type_definition = type_definition}; |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
884 |
|
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
885 |
val bnf_map = fold_rev Term.absfree fs' (HOLogic.mk_comp (HOLogic.mk_comp (AbsB, |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
886 |
Term.list_comb (mk_map_of_bnf Ds As Bs bnf, fs)), RepA)); |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
887 |
val bnf_sets = map ((fn t => HOLogic.mk_comp (t, RepA))) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
888 |
(mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf); |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
889 |
val bnf_bd = mk_bd_of_bnf Ds As bnf; |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
890 |
val bnf_rel = fold_rev Term.absfree Rs' (mk_vimage2p RepA RepB $ |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
891 |
(Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, Rs))); |
62324 | 892 |
val bnf_pred = fold_rev Term.absfree Ps' (HOLogic.mk_comp |
893 |
(Term.list_comb (mk_pred_of_bnf Ds As bnf, Ps), RepA)); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
894 |
|
55704
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
blanchet
parents:
55703
diff
changeset
|
895 |
(*bd may depend only on dead type variables*) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
896 |
val bd_repT = fst (dest_relT (fastype_of bnf_bd)); |
53264 | 897 |
val bdT_bind = qualify (Binding.suffix_name ("_" ^ bdTN) b); |
55707
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
traytel
parents:
55706
diff
changeset
|
898 |
val params = Term.add_tfreesT bd_repT []; |
63837
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
blanchet
parents:
63836
diff
changeset
|
899 |
val all_deads = map TFree (fold_rev Term.add_tfreesT all_Ds []); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
900 |
|
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
901 |
val ((bdT, (_, Abs_bd_name, _, _, Abs_bdT_inject, Abs_bdT_cases)), lthy) = |
63802 | 902 |
maybe_typedef false (bdT_bind, params, NoSyn) (HOLogic.mk_UNIV bd_repT) NONE |
903 |
(fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
904 |
|
75624 | 905 |
val (bnf_bd', bd_ordIso, bd_card_order, bd_cinfinite, bd_regularCard) = |
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
906 |
if bdT = bd_repT then (bnf_bd, bd_Card_order_of_bnf bnf RS @{thm ordIso_refl}, |
75624 | 907 |
bd_card_order_of_bnf bnf, bd_cinfinite_of_bnf bnf, bd_regularCard_of_bnf bnf) |
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
908 |
else |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
909 |
let |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
910 |
val bnf_bd' = mk_dir_image bnf_bd (Const (Abs_bd_name, bd_repT --> bdT)); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
911 |
|
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
912 |
val Abs_bdT_inj = mk_Abs_inj_thm Abs_bdT_inject; |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
913 |
val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj Abs_bdT_cases; |
57567 | 914 |
|
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
915 |
val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf]; |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
916 |
val bd_card_order = |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
917 |
@{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf]; |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
918 |
val bd_cinfinite = |
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
919 |
(@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1; |
75624 | 920 |
val bd_regularCard = @{thm regularCard_ordIso} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf, |
921 |
bd_regularCard_of_bnf bnf]; |
|
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
922 |
in |
75624 | 923 |
(bnf_bd', bd_ordIso, bd_card_order, bd_cinfinite, bd_regularCard) |
56012
158dc03db8be
made typedef for the type of the bound optional (size-based)
traytel
parents:
56010
diff
changeset
|
924 |
end; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
925 |
|
60728 | 926 |
fun map_id0_tac ctxt = |
927 |
rtac ctxt (@{thm type_copy_map_id0} OF [type_definition, map_id0_of_bnf bnf]) 1; |
|
928 |
fun map_comp0_tac ctxt = |
|
929 |
rtac ctxt (@{thm type_copy_map_comp0} OF [type_definition, map_comp0_of_bnf bnf]) 1; |
|
930 |
fun map_cong0_tac ctxt = |
|
931 |
EVERY' (rtac ctxt @{thm type_copy_map_cong0} :: rtac ctxt (map_cong0_of_bnf bnf) :: |
|
932 |
map (fn i => EVERY' [select_prem_tac ctxt live (dtac ctxt meta_spec) i, etac ctxt meta_mp, |
|
933 |
etac ctxt (o_apply RS equalityD2 RS set_mp)]) (1 upto live)) 1; |
|
934 |
fun set_map0_tac thm ctxt = |
|
935 |
rtac ctxt (@{thm type_copy_set_map0} OF [type_definition, thm]) 1; |
|
75624 | 936 |
val set_bd_tacs = map (fn thm => fn ctxt => rtac ctxt (@{thm ordLess_ordIso_trans} OF |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
937 |
[thm, bd_ordIso] RS @{thm type_copy_set_bd}) 1) (set_bd_of_bnf bnf); |
60728 | 938 |
fun le_rel_OO_tac ctxt = |
939 |
rtac ctxt (le_rel_OO_of_bnf bnf RS @{thm vimage2p_relcompp_mono}) 1; |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
940 |
fun rel_OO_Grp_tac ctxt = |
60728 | 941 |
(rtac ctxt (rel_OO_Grp_of_bnf bnf RS @{thm vimage2p_cong} RS trans) THEN' |
63802 | 942 |
(if force_out_of_line then subst_tac ctxt NONE else SELECT_GOAL o unfold_thms_tac ctxt) |
58359
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents:
58332
diff
changeset
|
943 |
[type_definition RS @{thm vimage2p_relcompp_converse}] THEN' |
58332 | 944 |
SELECT_GOAL (unfold_thms_tac ctxt [o_apply, |
945 |
type_definition RS @{thm type_copy_vimage2p_Grp_Rep}, |
|
946 |
type_definition RS @{thm vimage2p_relcompp_converse}]) THEN' |
|
60728 | 947 |
rtac ctxt refl) 1; |
62324 | 948 |
fun pred_set_tac ctxt = |
949 |
HEADGOAL (EVERY' |
|
67091 | 950 |
[rtac ctxt (pred_set_of_bnf bnf RS @{thm arg_cong[of _ _ "\<lambda>f. f \<circ> _"]} RS trans), |
62324 | 951 |
SELECT_GOAL (unfold_thms_tac ctxt (@{thms Ball_comp_iff conj_comp_iff})), rtac ctxt refl]); |
49456 | 952 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
953 |
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac |
60728 | 954 |
(map set_map0_tac (set_map0_of_bnf bnf)) |
955 |
(fn ctxt => rtac ctxt bd_card_order 1) (fn ctxt => rtac ctxt bd_cinfinite 1) |
|
75624 | 956 |
(fn ctxt => rtac ctxt bd_regularCard 1) set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac pred_set_tac; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
957 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
958 |
val bnf_wits = map (fn (I, t) => |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
959 |
fold Term.absdummy (map (nth As) I) |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
960 |
(AbsA $ Term.list_comb (t, map Bound (0 upto length I - 1)))) |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55707
diff
changeset
|
961 |
(mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
962 |
|
60752 | 963 |
fun wit_tac ctxt = |
964 |
ALLGOALS (dtac ctxt (type_definition RS @{thm type_copy_wit})) THEN |
|
965 |
mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
966 |
|
53264 | 967 |
val (bnf', lthy') = |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
968 |
bnf_def Hardly_Inline (user_policy Dont_Note) true qualify tacs wit_tac (SOME all_deads) |
62324 | 969 |
Binding.empty Binding.empty Binding.empty [] |
970 |
(((((((b, TA), bnf_map), bnf_sets), bnf_bd'), bnf_wits), SOME bnf_rel), SOME bnf_pred) lthy; |
|
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
971 |
|
58181 | 972 |
val unfolds = @{thm id_bnf_apply} :: |
71261 | 973 |
(#map_unfolds unfold_set @ flat (#set_unfoldss unfold_set) @ |
974 |
#rel_unfolds unfold_set @ #pred_unfolds unfold_set); |
|
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
975 |
|
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
976 |
val bnf'' = bnf' |> morph_bnf_defs (Morphism.thm_morphism "BNF" (unfold_thms lthy' unfolds)); |
57567 | 977 |
|
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
978 |
val map_def = map_def_of_bnf bnf''; |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
979 |
val set_defs = set_defs_of_bnf bnf''; |
56018 | 980 |
val rel_def = rel_def_of_bnf bnf''; |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
981 |
|
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
982 |
val bnf_b = qualify b; |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
983 |
val def_qualify = |
59859 | 984 |
Thm.def_binding o Binding.concealed o Binding.qualify false (Binding.name_of bnf_b); |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
985 |
fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b; |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
986 |
val map_b = def_qualify (mk_prefix_binding mapN); |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
987 |
val rel_b = def_qualify (mk_prefix_binding relN); |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
988 |
val set_bs = if live = 1 then [def_qualify (mk_prefix_binding setN)] |
59725 | 989 |
else map (def_qualify o mk_prefix_binding o mk_setN) (1 upto live); |
56016
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
990 |
|
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
991 |
val notes = (map_b, map_def) :: (rel_b, rel_def) :: (set_bs ~~ set_defs) |
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
traytel
parents:
56013
diff
changeset
|
992 |
|> map (fn (b, def) => ((b, []), [([def], [])])) |
57632 | 993 |
|
59820
0e9a0a5f4424
register pre-fixpoint BNFs in database to enable lookup later (e.g. in 'corec')
blanchet
parents:
59794
diff
changeset
|
994 |
val (noted, lthy'') = lthy' |
0e9a0a5f4424
register pre-fixpoint BNFs in database to enable lookup later (e.g. in 'corec')
blanchet
parents:
59794
diff
changeset
|
995 |
|> Local_Theory.notes notes |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
75624
diff
changeset
|
996 |
||> (if repTA = TA then I else register_bnf_raw (dest_Type_name TA) bnf'') |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
997 |
in |
57632 | 998 |
((morph_bnf (substitute_noted_thm noted) bnf'', (all_deads, absT_info)), lthy'') |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
999 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1000 |
|
53222 | 1001 |
exception BAD_DEAD of typ * typ; |
1002 |
||
62621 | 1003 |
fun bnf_of_typ _ _ _ _ _ Ds0 (T as TFree T') accum = |
55704
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
blanchet
parents:
55703
diff
changeset
|
1004 |
(if member (op =) Ds0 T' then (DEADID_bnf, ([T], [])) else (ID_bnf, ([], [T])), accum) |
62621 | 1005 |
| bnf_of_typ _ _ _ _ _ _ (TVar _) _ = error "Unexpected schematic variable" |
1006 |
| bnf_of_typ optim const_policy qualify' flatten_tyargs Xs Ds0 (T as Type (C, Ts)) |
|
58332 | 1007 |
(accum as (_, lthy)) = |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1008 |
let |
53222 | 1009 |
fun check_bad_dead ((_, (deads, _)), _) = |
1010 |
let val Ds = fold Term.add_tfreesT deads [] in |
|
1011 |
(case Library.inter (op =) Ds Xs of [] => () |
|
55705 | 1012 |
| X :: _ => raise BAD_DEAD (TFree X, T)) |
53222 | 1013 |
end; |
1014 |
||
55704
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
blanchet
parents:
55703
diff
changeset
|
1015 |
val tfrees = subtract (op =) Ds0 (Term.add_tfreesT T []); |
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
blanchet
parents:
55703
diff
changeset
|
1016 |
val bnf_opt = if null tfrees then NONE else bnf_of lthy C; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1017 |
in |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1018 |
(case bnf_opt of |
55704
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
blanchet
parents:
55703
diff
changeset
|
1019 |
NONE => ((DEADID_bnf, ([T], [])), accum) |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1020 |
| SOME bnf => |
62621 | 1021 |
if optim andalso forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1022 |
let |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1023 |
val T' = T_of_bnf bnf; |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1024 |
val deads = deads_of_bnf bnf; |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1025 |
val lives = lives_of_bnf bnf; |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1026 |
val tvars' = Term.add_tvarsT T' []; |
55904 | 1027 |
val Ds_As = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58959
diff
changeset
|
1028 |
apply2 (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees))) |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1029 |
(deads, lives); |
55904 | 1030 |
in ((bnf, Ds_As), accum) end |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1031 |
else |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1032 |
let |
49425 | 1033 |
val name = Long_Name.base_name C; |
1034 |
fun qualify i = |
|
1035 |
let val namei = name ^ nonzero_string_of_int i; |
|
1036 |
in qualify' o Binding.qualify true namei end; |
|
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1037 |
val odead = dead_of_bnf bnf; |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1038 |
val olive = live_of_bnf bnf; |
59131
894d613f7f7c
respect order of deads when retrieving bnfs from the database
traytel
parents:
59058
diff
changeset
|
1039 |
val Ds = map (fn i => TFree (string_of_int i, [])) (1 upto odead); |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
75624
diff
changeset
|
1040 |
val Us = Term.dest_Type_args (mk_T_of_bnf Ds (replicate olive dummyT) bnf); |
59131
894d613f7f7c
respect order of deads when retrieving bnfs from the database
traytel
parents:
59058
diff
changeset
|
1041 |
val oDs_pos = map (fn x => find_index (fn y => x = y) Us) Ds |
894d613f7f7c
respect order of deads when retrieving bnfs from the database
traytel
parents:
59058
diff
changeset
|
1042 |
|> filter (fn x => x >= 0); |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1043 |
val oDs = map (nth Ts) oDs_pos; |
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1044 |
val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1)); |
55904 | 1045 |
val ((inners, (Dss, Ass)), (accum', lthy')) = |
62621 | 1046 |
apfst (apsnd split_list o split_list) (@{fold_map 2} |
1047 |
(fn i => bnf_of_typ optim Smart_Inline (qualify i) flatten_tyargs Xs Ds0) |
|
59725 | 1048 |
(if length Ts' = 1 then [0] else 1 upto length Ts') Ts' accum); |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1049 |
in |
58332 | 1050 |
compose_bnf const_policy qualify flatten_tyargs bnf inners oDs Dss Ass (accum', lthy') |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
1051 |
end) |
53222 | 1052 |
|> tap check_bad_dead |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1053 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1054 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1055 |
end; |