| author | Thomas Lindae <thomas.lindae@in.tum.de> | 
| Fri, 14 Jun 2024 10:21:03 +0200 | |
| changeset 81054 | 4bfcb14547c6 | 
| parent 80636 | 4041e7c8059d | 
| child 82630 | 2bb4a8d0111d | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_lfp.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: Andrei Popescu, TU Muenchen  | 
| 75624 | 4  | 
Author: Jan van Brügge, TU Muenchen  | 
5  | 
Copyright 2012, 2022  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
|
| 58315 | 7  | 
Datatype construction.  | 
| 
48975
 
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  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
signature BNF_LFP =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
11  | 
sig  | 
| 62324 | 12  | 
val construct_lfp: mixfix list -> binding list -> binding list -> binding list ->  | 
13  | 
binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->  | 
|
14  | 
BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory ->  | 
|
15  | 
BNF_FP_Util.fp_result * local_theory  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
16  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
18  | 
structure BNF_LFP : BNF_LFP =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
19  | 
struct  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
21  | 
open BNF_Def  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
22  | 
open BNF_Util  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
23  | 
open BNF_Tactics  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
24  | 
open BNF_Comp  | 
| 
51850
 
106afdf5806c
renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
 
blanchet 
parents: 
51839 
diff
changeset
 | 
25  | 
open BNF_FP_Util  | 
| 49636 | 26  | 
open BNF_FP_Def_Sugar  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
27  | 
open BNF_LFP_Util  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
28  | 
open BNF_LFP_Tactics  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
29  | 
|
| 
49460
 
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
 
blanchet 
parents: 
49459 
diff
changeset
 | 
30  | 
(*all BNFs have the same lives*)  | 
| 
62684
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62324 
diff
changeset
 | 
31  | 
fun construct_lfp mixfixes map_bs rel_bs pred_bs set_bss0 bs resBs (resDs, Dss) bnfs absT_infos  | 
| 
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62324 
diff
changeset
 | 
32  | 
lthy =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
33  | 
let  | 
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53138 
diff
changeset
 | 
34  | 
val time = time lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
35  | 
val timer = time (Timer.startRealTimer ());  | 
| 
49580
 
040cfb087b3a
leave out some internal theorems unless "bnf_note_all" is set
 
blanchet 
parents: 
49545 
diff
changeset
 | 
36  | 
|
| 49132 | 37  | 
val live = live_of_bnf (hd bnfs);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
38  | 
val n = length bnfs; (*active*)  | 
| 49132 | 39  | 
val ks = 1 upto n;  | 
| 
49460
 
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
 
blanchet 
parents: 
49459 
diff
changeset
 | 
40  | 
val m = live - n; (*passive, if 0 don't generate a new BNF*)  | 
| 53566 | 41  | 
|
| 62093 | 42  | 
val internals = Config.get lthy bnf_internals;  | 
| 52956 | 43  | 
val b_names = map Binding.name_of bs;  | 
| 53566 | 44  | 
val b_name = mk_common_name b_names;  | 
45  | 
val b = Binding.name b_name;  | 
|
| 
58241
 
ff8059e3e803
generate better internal names, with name of the target type in it
 
blanchet 
parents: 
58208 
diff
changeset
 | 
46  | 
|
| 
 
ff8059e3e803
generate better internal names, with name of the target type in it
 
blanchet 
parents: 
58208 
diff
changeset
 | 
47  | 
fun mk_internal_of_b name =  | 
| 59859 | 48  | 
Binding.prefix_name (name ^ "_") #> Binding.prefix true b_name #> Binding.concealed;  | 
| 
58241
 
ff8059e3e803
generate better internal names, with name of the target type in it
 
blanchet 
parents: 
58208 
diff
changeset
 | 
49  | 
fun mk_internal_b name = mk_internal_of_b name b;  | 
| 
 
ff8059e3e803
generate better internal names, with name of the target type in it
 
blanchet 
parents: 
58208 
diff
changeset
 | 
50  | 
fun mk_internal_bs name = map (mk_internal_of_b name) bs;  | 
| 53566 | 51  | 
val external_bs = map2 (Binding.prefix false) b_names bs  | 
| 62093 | 52  | 
|> not internals ? map Binding.concealed;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
53  | 
|
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
54  | 
val deads = fold (union (op =)) Dss resDs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
55  | 
val names_lthy = fold Variable.declare_typ deads lthy;  | 
| 52956 | 56  | 
val passives = map fst (subtract (op = o apsnd TFree) deads resBs);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
(* tvars *)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
59  | 
val (((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs) =  | 
| 
52938
 
3b3e52d5e66b
tuned name generation code (to make it easier to adapt later)
 
blanchet 
parents: 
52923 
diff
changeset
 | 
60  | 
names_lthy  | 
| 52956 | 61  | 
|> variant_tfrees passives  | 
| 
52938
 
3b3e52d5e66b
tuned name generation code (to make it easier to adapt later)
 
blanchet 
parents: 
52923 
diff
changeset
 | 
62  | 
||>> mk_TFrees n  | 
| 52956 | 63  | 
||>> variant_tfrees passives  | 
| 
52938
 
3b3e52d5e66b
tuned name generation code (to make it easier to adapt later)
 
blanchet 
parents: 
52923 
diff
changeset
 | 
64  | 
||>> mk_TFrees n  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
65  | 
||>> variant_tfrees passives  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
66  | 
||>> mk_TFrees n  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
67  | 
|> fst;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
68  | 
|
| 
52938
 
3b3e52d5e66b
tuned name generation code (to make it easier to adapt later)
 
blanchet 
parents: 
52923 
diff
changeset
 | 
69  | 
val allAs = passiveAs @ activeAs;  | 
| 
 
3b3e52d5e66b
tuned name generation code (to make it easier to adapt later)
 
blanchet 
parents: 
52923 
diff
changeset
 | 
70  | 
val allBs' = passiveBs @ activeBs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
71  | 
val Ass = replicate n allAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
72  | 
val allBs = passiveAs @ activeBs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
73  | 
val Bss = replicate n allBs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
74  | 
val allCs = passiveAs @ activeCs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
75  | 
val allCs' = passiveBs @ activeCs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
76  | 
val Css' = replicate n allCs';  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
77  | 
|
| 51866 | 78  | 
(* types *)  | 
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
79  | 
val dead_poss =  | 
| 52956 | 80  | 
map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;  | 
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
81  | 
fun mk_param NONE passive = (hd passive, tl passive)  | 
| 
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
82  | 
| mk_param (SOME a) passive = (a, passive);  | 
| 
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
83  | 
val mk_params = fold_map mk_param dead_poss #> fst;  | 
| 
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
84  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
85  | 
fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;  | 
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
86  | 
val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
87  | 
val FTsAs = mk_FTs allAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
88  | 
val FTsBs = mk_FTs allBs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
89  | 
val FTsCs = mk_FTs allCs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
90  | 
val BTs = map HOLogic.mk_setT activeAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
91  | 
val B'Ts = map HOLogic.mk_setT activeBs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
92  | 
val B''Ts = map HOLogic.mk_setT activeCs;  | 
| 52923 | 93  | 
val sTs = map2 (curry op -->) FTsAs activeAs;  | 
94  | 
val s'Ts = map2 (curry op -->) FTsBs activeBs;  | 
|
95  | 
val s''Ts = map2 (curry op -->) FTsCs activeCs;  | 
|
96  | 
val fTs = map2 (curry op -->) activeAs activeBs;  | 
|
97  | 
val inv_fTs = map2 (curry op -->) activeBs activeAs;  | 
|
98  | 
val self_fTs = map2 (curry op -->) activeAs activeAs;  | 
|
99  | 
val gTs = map2 (curry op -->) activeBs activeCs;  | 
|
100  | 
val all_gTs = map2 (curry op -->) allBs allCs';  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
102  | 
(* terms *)  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
103  | 
    val mapsAsAs = @{map 4} mk_map_of_bnf Dss Ass Ass bnfs;
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
104  | 
    val mapsAsBs = @{map 4} mk_map_of_bnf Dss Ass Bss bnfs;
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
105  | 
    val mapsBsCs' = @{map 4} mk_map_of_bnf Dss Bss Css' bnfs;
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
106  | 
    val mapsAsCs' = @{map 4} mk_map_of_bnf Dss Ass Css' bnfs;
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
107  | 
    fun mk_setss Ts = @{map 3} mk_sets_of_bnf (map (replicate live) Dss)
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
108  | 
(map (replicate live) (replicate n Ts)) bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
109  | 
val setssAs = mk_setss allAs;  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
110  | 
    val bd0s = @{map 3} mk_bd_of_bnf Dss Ass bnfs;
 | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
111  | 
val bds =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
112  | 
      @{map 3} (fn bd0 => fn Ds => fn bnf => mk_csum bd0
 | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
113  | 
(mk_card_of (HOLogic.mk_UNIV  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
114  | 
(mk_T_of_bnf Ds (replicate live (fst (dest_relT (fastype_of bd0)))) bnf))))  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
115  | 
bd0s Dss bnfs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
116  | 
val witss = map wits_of_bnf bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
117  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
118  | 
val ((((((((zs, zs'), Bs), ss), fs), self_fs), all_gs), (xFs, xFs')), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
119  | 
lthy  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
120  | 
|> mk_Frees' "z" activeAs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
121  | 
||>> mk_Frees "B" BTs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
122  | 
||>> mk_Frees "s" sTs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
123  | 
||>> mk_Frees "f" fTs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
124  | 
||>> mk_Frees "f" self_fTs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
125  | 
||>> mk_Frees "g" all_gTs  | 
| 55756 | 126  | 
||>> mk_Frees' "x" FTsAs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
127  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
128  | 
val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
129  | 
val active_UNIVs = map HOLogic.mk_UNIV activeAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
130  | 
val passive_ids = map HOLogic.id_const passiveAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
val active_ids = map HOLogic.id_const activeAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
132  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
133  | 
(* thms *)  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
134  | 
val bd0_card_orders = map bd_card_order_of_bnf bnfs;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
135  | 
val bd0_Card_orders = map bd_Card_order_of_bnf bnfs;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
136  | 
val bd0_Cinfinites = map bd_Cinfinite_of_bnf bnfs;  | 
| 75624 | 137  | 
val bd0_regularCards = map bd_regularCard_of_bnf bnfs;  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
138  | 
val set_bd0ss = map set_bd_of_bnf bnfs;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
139  | 
|
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
140  | 
    val bd_Card_order = @{thm Card_order_csum};
 | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
141  | 
val bd_Card_orders = replicate n bd_Card_order;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
142  | 
    val bd_Cinfinites = map (fn thm => thm RS @{thm Cinfinite_csum1}) bd0_Cinfinites;
 | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
143  | 
    val bd_Cnotzeros = map (fn thm => thm RS @{thm Cinfinite_Cnotzero}) bd_Cinfinites;
 | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
144  | 
val bd_Cinfinite = hd bd_Cinfinites;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
145  | 
val set_bdss =  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
146  | 
map2 (fn set_bd0s => fn bd0_Card_order =>  | 
| 75624 | 147  | 
        map (fn thm => @{thm ordLess_ordLeq_trans} OF [thm, bd0_Card_order RS @{thm ordLeq_csum1}]) set_bd0s)
 | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
52505 
diff
changeset
 | 
148  | 
set_bd0ss bd0_Card_orders;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
149  | 
val in_bds = map in_bd_of_bnf bnfs;  | 
| 53287 | 150  | 
val sym_map_comps = map (fn bnf => map_comp0_of_bnf bnf RS sym) bnfs;  | 
| 53288 | 151  | 
val map_comps = map map_comp_of_bnf bnfs;  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
152  | 
val map_cong0s = map map_cong0_of_bnf bnfs;  | 
| 53270 | 153  | 
val map_id0s = map map_id0_of_bnf bnfs;  | 
| 53285 | 154  | 
val map_ids = map map_id_of_bnf bnfs;  | 
| 53290 | 155  | 
val set_mapss = map set_map_of_bnf bnfs;  | 
| 57967 | 156  | 
val rel_mono_strong0s = map rel_mono_strong0_of_bnf bnfs;  | 
| 57726 | 157  | 
val le_rel_OOs = map le_rel_OO_of_bnf bnfs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
val timer = time (timer "Extracted terms & thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
161  | 
(* nonemptiness check *)  | 
| 51070 | 162  | 
fun new_wit X (wit: nonemptiness_witness) = subset (op =) (#I wit, (0 upto m - 1) @ map snd X);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
163  | 
|
| 
49341
 
d406979024d1
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
 
blanchet 
parents: 
49337 
diff
changeset
 | 
164  | 
val all = m upto m + n - 1;  | 
| 
 
d406979024d1
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
 
blanchet 
parents: 
49337 
diff
changeset
 | 
165  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
166  | 
fun enrich X = map_filter (fn i =>  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
167  | 
(case find_first (fn (_, i') => i = i') X of  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
168  | 
NONE =>  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
169  | 
(case find_index (new_wit X) (nth witss (i - m)) of  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
170  | 
~1 => NONE  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
171  | 
| j => SOME (j, i))  | 
| 
49341
 
d406979024d1
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
 
blanchet 
parents: 
49337 
diff
changeset
 | 
172  | 
| SOME ji => SOME ji)) all;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
173  | 
val reachable = fixpoint (op =) enrich [];  | 
| 
49341
 
d406979024d1
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
 
blanchet 
parents: 
49337 
diff
changeset
 | 
174  | 
val _ = (case subtract (op =) (map snd reachable) all of  | 
| 
 
d406979024d1
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
 
blanchet 
parents: 
49337 
diff
changeset
 | 
175  | 
[] => ()  | 
| 
63826
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63312 
diff
changeset
 | 
176  | 
| i :: _ => raise EMPTY_DATATYPE (Binding.name_of (nth bs (i - m))));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
177  | 
|
| 
49341
 
d406979024d1
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
 
blanchet 
parents: 
49337 
diff
changeset
 | 
178  | 
val wit_thms = flat (map2 (fn bnf => fn (j, _) => nth (wit_thmss_of_bnf bnf) j) bnfs reachable);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
180  | 
val timer = time (timer "Checked nonemptiness");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
182  | 
(* derived thms *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
183  | 
|
| 52956 | 184  | 
(*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
185  | 
map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)  | 
| 53287 | 186  | 
fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
187  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
188  | 
val lhs = Term.list_comb (mapBsCs, all_gs) $  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
189  | 
(Term.list_comb (mapAsBs, passive_ids @ fs) $ x);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
190  | 
val rhs = Term.list_comb (mapAsCs,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
191  | 
take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
192  | 
val vars = fold (Variable.add_free_names lthy) [lhs, rhs] [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
193  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
194  | 
Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs))  | 
| 55756 | 195  | 
          (fn {context = ctxt, prems = _} => mk_map_comp_id_tac ctxt map_comp0)
 | 
| 70494 | 196  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
197  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
198  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
199  | 
    val map_comp_id_thms = @{map 5} mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
200  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
201  | 
(*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
202  | 
map id ... id f(m+1) ... f(m+n) x = x*)  | 
| 53285 | 203  | 
fun mk_map_cong0L x mapAsAs sets map_cong0 map_id =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
204  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
205  | 
fun mk_prem set f z z' = HOLogic.mk_Trueprop  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
206  | 
(mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
207  | 
        val prems = @{map 4} mk_prem (drop m sets) self_fs zs zs';
 | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents: 
49121 
diff
changeset
 | 
208  | 
val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
209  | 
val vars = fold (Variable.add_free_names lthy) (goal :: prems) [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
210  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
211  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal))  | 
| 60728 | 212  | 
          (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id)
 | 
| 70494 | 213  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
214  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
215  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
216  | 
    val map_cong0L_thms = @{map 5} mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
 | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
217  | 
val in_mono'_thms = map (fn bnf => in_mono_of_bnf bnf OF (replicate m subset_refl)) bnfs;  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
218  | 
val in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
220  | 
val timer = time (timer "Derived simple theorems");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
222  | 
(* algebra *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
223  | 
|
| 53566 | 224  | 
val alg_bind = mk_internal_b algN;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
225  | 
val alg_def_bind = (Thm.def_binding alg_bind, []);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
226  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
227  | 
(*forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV .. UNIV B1 ... Bn. si x \<in> Bi)*)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
228  | 
val alg_spec =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
229  | 
let  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
230  | 
        val ins = @{map 3} mk_in (replicate n (passive_UNIVs @ Bs)) setssAs FTsAs;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
231  | 
fun mk_alg_conjunct B s X x x' =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
232  | 
mk_Ball X (Term.absfree x' (HOLogic.mk_mem (s $ x, B)));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
233  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
234  | 
        val rhs = Library.foldr1 HOLogic.mk_conj (@{map 5} mk_alg_conjunct Bs ss ins xFs xFs')
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
235  | 
in  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
236  | 
fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) rhs  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
237  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
239  | 
val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) =  | 
| 58332 | 240  | 
lthy  | 
| 72450 | 241  | 
|> (snd o Local_Theory.begin_nested)  | 
| 58332 | 242  | 
|> Local_Theory.define ((alg_bind, NoSyn), (alg_def_bind, alg_spec))  | 
| 72450 | 243  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
245  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
246  | 
val alg = dest_Const_name (Morphism.term phi alg_free);  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
247  | 
val alg_def = mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi alg_def_free));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
248  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
249  | 
fun mk_alg Bs ss =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
250  | 
let  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
251  | 
val args = Bs @ ss;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
252  | 
val Ts = map fastype_of args;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
253  | 
val algT = Library.foldr (op -->) (Ts, HOLogic.boolT);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
254  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
255  | 
Term.list_comb (Const (alg, algT), args)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
256  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
257  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
258  | 
val ((((((((zs, zs'), Bs), B's), ss), s's), fs), (xFs, xFs')), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
259  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
260  | 
|> mk_Frees' "z" activeAs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
261  | 
||>> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
262  | 
||>> mk_Frees "B'" B'Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
263  | 
||>> mk_Frees "s" sTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
264  | 
||>> mk_Frees "s'" s'Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
265  | 
||>> mk_Frees "f" fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
266  | 
||>> mk_Frees' "x" FTsAs;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
267  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
268  | 
val alg_set_thms =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
269  | 
let  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
270  | 
val alg_prem = HOLogic.mk_Trueprop (mk_alg Bs ss);  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51869 
diff
changeset
 | 
271  | 
fun mk_prem x set B = HOLogic.mk_Trueprop (mk_leq (set $ x) B);  | 
| 57567 | 272  | 
fun mk_concl s x B = mk_Trueprop_mem (s $ x, B);  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
273  | 
val premss = map2 ((fn x => fn sets => map2 (mk_prem x) (drop m sets) Bs)) xFs setssAs;  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
274  | 
        val concls = @{map 3} mk_concl ss xFs Bs;
 | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
275  | 
val goals = map2 (fn prems => fn concl =>  | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
276  | 
Logic.list_implies (alg_prem :: prems, concl)) premss concls;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
277  | 
in  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
278  | 
map (fn goal =>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
279  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
280  | 
          |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
281  | 
mk_alg_set_tac ctxt alg_def))  | 
| 70494 | 282  | 
|> Thm.close_derivation \<^here>)  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
283  | 
goals  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
284  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
285  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
286  | 
val timer = time (timer "Algebra definition & thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
287  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
288  | 
val alg_not_empty_thms =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
289  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
290  | 
val alg_prem =  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
291  | 
HOLogic.mk_Trueprop (mk_alg Bs ss);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
292  | 
val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
293  | 
val goals =  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
294  | 
map (fn concl => Logic.mk_implies (alg_prem, concl)) concls;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
295  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
296  | 
map2 (fn goal => fn alg_set =>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
297  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
298  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 61271 | 299  | 
            (fn {context = ctxt, prems = _} =>
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
300  | 
mk_alg_not_empty_tac ctxt alg_set alg_set_thms wit_thms))  | 
| 70494 | 301  | 
|> Thm.close_derivation \<^here>)  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
302  | 
goals alg_set_thms  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
303  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
304  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
305  | 
val timer = time (timer "Proved nonemptiness");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
307  | 
(* morphism *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
308  | 
|
| 53566 | 309  | 
val mor_bind = mk_internal_b morN;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
310  | 
val mor_def_bind = (Thm.def_binding mor_bind, []);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
312  | 
(*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. f x \<in> B'i)*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
313  | 
(*mor) forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV ... UNIV B1 ... Bn.  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
314  | 
f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
315  | 
val mor_spec =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
316  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
317  | 
fun mk_fbetw f B1 B2 z z' =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
318  | 
mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
319  | 
fun mk_mor sets mapAsBs f s s' T x x' =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
320  | 
mk_Ball (mk_in (passive_UNIVs @ Bs) sets T)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
321  | 
(Term.absfree x' (HOLogic.mk_eq (f $ (s $ x), s' $  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
322  | 
(Term.list_comb (mapAsBs, passive_ids @ fs) $ x))));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
323  | 
val rhs = HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
324  | 
          (Library.foldr1 HOLogic.mk_conj (@{map 5} mk_fbetw fs Bs B's zs zs'),
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
325  | 
Library.foldr1 HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
326  | 
            (@{map 8} mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs'))
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
327  | 
in  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
328  | 
fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ fs) rhs  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
329  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
330  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
331  | 
val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =  | 
| 59794 | 332  | 
lthy  | 
| 72450 | 333  | 
|> (snd o Local_Theory.begin_nested)  | 
| 59794 | 334  | 
|> Local_Theory.define ((mor_bind, NoSyn), (mor_def_bind, mor_spec))  | 
| 72450 | 335  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
336  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
337  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
338  | 
val mor = dest_Const_name (Morphism.term phi mor_free);  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
339  | 
val mor_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi mor_def_free));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
340  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
341  | 
fun mk_mor Bs1 ss1 Bs2 ss2 fs =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
342  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
343  | 
val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
344  | 
val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
345  | 
val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
346  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
347  | 
Term.list_comb (Const (mor, morT), args)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
348  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
349  | 
|
| 63312 | 350  | 
val (((((((((((Bs, Bs_copy), B's), B''s), ss), s's), s''s), fs), fs_copy), gs), xFs), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
351  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
352  | 
|> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
353  | 
||>> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
354  | 
||>> mk_Frees "B'" B'Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
355  | 
||>> mk_Frees "B''" B''Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
356  | 
||>> mk_Frees "s" sTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
357  | 
||>> mk_Frees "s'" s'Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
358  | 
||>> mk_Frees "s''" s''Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
359  | 
||>> mk_Frees "f" fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
360  | 
||>> mk_Frees "f" fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
361  | 
||>> mk_Frees "g" gTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
362  | 
||>> mk_Frees "x" FTsAs;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
363  | 
|
| 56237 | 364  | 
val morE_thms =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
365  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
366  | 
val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
367  | 
fun mk_elim_prem sets x T = HOLogic.mk_Trueprop  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
368  | 
(HOLogic.mk_mem (x, mk_in (passive_UNIVs @ Bs) sets T));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
369  | 
fun mk_elim_goal sets mapAsBs f s s' x T =  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
370  | 
Logic.list_implies ([prem, mk_elim_prem sets x T],  | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
371  | 
mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x])));  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
372  | 
        val elim_goals = @{map 7} mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs;
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
373  | 
fun prove goal =  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
374  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
375  | 
          |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
376  | 
mk_mor_elim_tac ctxt mor_def))  | 
| 70494 | 377  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
378  | 
in  | 
| 56237 | 379  | 
map prove elim_goals  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
380  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
381  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
382  | 
val mor_incl_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
383  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51869 
diff
changeset
 | 
384  | 
val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
385  | 
val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
386  | 
val vars = fold (Variable.add_free_names lthy) (concl :: prems) [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
387  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
388  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl))  | 
| 60728 | 389  | 
          (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids)
 | 
| 70494 | 390  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
391  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
392  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
393  | 
val mor_comp_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
394  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
395  | 
val prems =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
396  | 
[HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
397  | 
HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
398  | 
val concl =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
399  | 
HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
400  | 
val vars = fold (Variable.add_free_names lthy) (concl :: prems) [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
401  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
402  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl))  | 
| 60728 | 403  | 
          (fn {context = ctxt, prems = _} => mk_mor_comp_tac ctxt mor_def set_mapss map_comp_id_thms)
 | 
| 70494 | 404  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
405  | 
end;  | 
| 
 
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 mor_cong_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
408  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
409  | 
val prems = map HOLogic.mk_Trueprop  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
410  | 
(map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
411  | 
val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
412  | 
val vars = fold (Variable.add_free_names lthy) (concl :: prems) [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
413  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
414  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl))  | 
| 60728 | 415  | 
          (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1)
 | 
| 70494 | 416  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
417  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
418  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
419  | 
val mor_str_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
420  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
421  | 
val maps = map2 (fn Ds => fn bnf => Term.list_comb  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
422  | 
(mk_map_of_bnf Ds (passiveAs @ FTsAs) allAs bnf, passive_ids @ ss)) Dss bnfs;  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
423  | 
val goal = HOLogic.mk_Trueprop  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
424  | 
(mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss);  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
425  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
426  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
427  | 
Goal.prove_sorry lthy vars [] goal  | 
| 60728 | 428  | 
          (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_def)
 | 
| 70494 | 429  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
430  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
431  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
432  | 
val mor_UNIV_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
433  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
434  | 
fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
435  | 
(HOLogic.mk_comp (f, s),  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
436  | 
HOLogic.mk_comp (s', Term.list_comb (mapAsBs, passive_ids @ fs)));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
437  | 
val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
438  | 
        val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's);
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
439  | 
val vars = fold (Variable.add_free_names lthy) [lhs, rhs] [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
440  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
441  | 
Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs))  | 
| 60728 | 442  | 
          (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt m morE_thms mor_def)
 | 
| 70494 | 443  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
444  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
445  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
446  | 
val timer = time (timer "Morphism definition & thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
447  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
448  | 
(* bounds *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
449  | 
|
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
450  | 
val sum_bd = Library.foldr1 (uncurry mk_csum) bds;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
451  | 
val sum_bdT = fst (dest_relT (fastype_of sum_bd));  | 
| 
56516
 
a13c2ccc160b
more accurate type arguments for intermeadiate typedefs
 
traytel 
parents: 
56350 
diff
changeset
 | 
452  | 
val (sum_bdT_params, sum_bdT_params') = `(map TFree) (Term.add_tfreesT sum_bdT []);  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
453  | 
|
| 56350 | 454  | 
val (lthy, sbd, sbd_Cinfinite, sbd_Card_order, set_sbdss, in_sbds) =  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
455  | 
if n = 1  | 
| 56350 | 456  | 
then (lthy, sum_bd, bd_Cinfinite, bd_Card_order, set_bdss, in_bds)  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
457  | 
else  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
458  | 
let  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
459  | 
val sbdT_bind = mk_internal_b sum_bdTN;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
460  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
461  | 
val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =  | 
| 
56516
 
a13c2ccc160b
more accurate type arguments for intermeadiate typedefs
 
traytel 
parents: 
56350 
diff
changeset
 | 
462  | 
typedef (sbdT_bind, sum_bdT_params', NoSyn)  | 
| 60728 | 463  | 
(HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt =>  | 
464  | 
EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy;  | 
|
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
465  | 
|
| 
56516
 
a13c2ccc160b
more accurate type arguments for intermeadiate typedefs
 
traytel 
parents: 
56350 
diff
changeset
 | 
466  | 
val sbdT = Type (sbdT_name, sum_bdT_params);  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
467  | 
val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT);  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
468  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
469  | 
val sbd_bind = mk_internal_b sum_bdN;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
470  | 
val sbd_def_bind = (Thm.def_binding sbd_bind, []);  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
471  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
472  | 
val sbd_spec = mk_dir_image sum_bd Abs_sbdT;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
473  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
474  | 
val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) =  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
475  | 
lthy  | 
| 72450 | 476  | 
|> (snd o Local_Theory.begin_nested)  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
477  | 
|> Local_Theory.define ((sbd_bind, NoSyn), (sbd_def_bind, sbd_spec))  | 
| 72450 | 478  | 
||> `Local_Theory.end_nested;  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
479  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
480  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
481  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
482  | 
val sbd_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd_def_free);  | 
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
483  | 
val sbd = Const (dest_Const_name (Morphism.term phi sbd_free), mk_relT (`I sbdT));  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
484  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
485  | 
val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
486  | 
|
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
487  | 
val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
488  | 
val sum_Card_order = sum_Cinfinite RS conjunct2;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
489  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
490  | 
          val sbd_ordIso = @{thm ssubst_Pair_rhs} OF
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
491  | 
            [@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order], sbd_def];
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
492  | 
          val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite];
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
493  | 
val sbd_Card_order = sbd_Cinfinite RS conjunct2;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
494  | 
|
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
495  | 
fun mk_set_sbd i bd_Card_order bds =  | 
| 75624 | 496  | 
            map (fn thm => @{thm ordLess_ordIso_trans} OF
 | 
497  | 
[bd_Card_order RS mk_ordLess_csum n i thm, sbd_ordIso]) bds;  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
498  | 
          val set_sbdss = @{map 3} mk_set_sbd ks bd_Card_orders set_bdss;
 | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
499  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
500  | 
fun mk_in_bd_sum i Co Cnz bd =  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
501  | 
            Cnz RS ((@{thm ordLeq_ordIso_trans} OF
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
502  | 
              [Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl}), sbd_ordIso]) RS
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
503  | 
              (bd RS @{thm ordLeq_transitive[OF _ cexp_mono2_Cnotzero[OF _ Card_order_csum]]}));
 | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
504  | 
          val in_sbds = @{map 4} mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds;
 | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
505  | 
in  | 
| 56350 | 506  | 
(lthy, sbd, sbd_Cinfinite, sbd_Card_order, set_sbdss, in_sbds)  | 
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
507  | 
end;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
508  | 
|
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
509  | 
    val sbd_Cnotzero = sbd_Cinfinite RS @{thm Cinfinite_Cnotzero};
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
510  | 
val suc_bd = mk_cardSuc sbd;  | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
511  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
512  | 
val field_suc_bd = mk_Field suc_bd;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
513  | 
val suc_bdT = fst (dest_relT (fastype_of suc_bd));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
514  | 
fun mk_Asuc_bd [] = mk_cexp ctwo suc_bd  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
515  | 
| mk_Asuc_bd As =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
516  | 
mk_cexp (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) suc_bd;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
517  | 
|
| 
55770
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
518  | 
    val suc_bd_Card_order =  sbd_Card_order RS @{thm cardSuc_Card_order};
 | 
| 
 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 
traytel 
parents: 
55756 
diff
changeset
 | 
519  | 
    val suc_bd_Cinfinite = sbd_Cinfinite RS @{thm Cinfinite_cardSuc};
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
520  | 
    val suc_bd_Cnotzero = suc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
521  | 
    val suc_bd_worel = suc_bd_Card_order RS @{thm Card_order_wo_rel}
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
522  | 
    val basis_Asuc = if m = 0 then @{thm ordLeq_refl[OF Card_order_ctwo]}
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
523  | 
        else @{thm ordLeq_csum2[OF Card_order_ctwo]};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
524  | 
    val Asuc_bd_Cinfinite = suc_bd_Cinfinite RS (basis_Asuc RS @{thm Cinfinite_cexp});
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
525  | 
|
| 
51782
 
84e7225f5ab6
removed unnecessary assumptions in some theorems about cardinal exponentiation
 
traytel 
parents: 
51767 
diff
changeset
 | 
526  | 
    val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF ordLess_ctwo_cexp cexp_mono1]} OF
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
527  | 
[suc_bd_Card_order, basis_Asuc, suc_bd_Card_order];  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
528  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
529  | 
|
| 
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
530  | 
val Asuc_bd = mk_Asuc_bd passive_UNIVs;  | 
| 
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
531  | 
val Asuc_bdT = fst (dest_relT (fastype_of Asuc_bd));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
532  | 
val II_BTs = replicate n (HOLogic.mk_setT Asuc_bdT);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
533  | 
val II_sTs = map2 (fn Ds => fn bnf =>  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
534  | 
mk_T_of_bnf Ds (passiveAs @ replicate n Asuc_bdT) bnf --> Asuc_bdT) Dss bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
535  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
536  | 
val ((((((Bs, ss), idxs), Asi_name), (idx, idx')), (jdx, jdx')), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
537  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
538  | 
|> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
539  | 
||>> mk_Frees "s" sTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
540  | 
||>> mk_Frees "i" (replicate n suc_bdT)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
541  | 
||>> (fn ctxt => apfst the_single (mk_fresh_names ctxt 1 "Asi"))  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
542  | 
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") suc_bdT  | 
| 
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
543  | 
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "j") suc_bdT;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
544  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
545  | 
val suc_bd_limit_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
546  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
547  | 
val prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
548  | 
(map (fn idx => HOLogic.mk_mem (idx, field_suc_bd)) idxs));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
549  | 
fun mk_conjunct idx = HOLogic.mk_conj (mk_not_eq idx jdx,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
550  | 
HOLogic.mk_mem (HOLogic.mk_prod (idx, jdx), suc_bd));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
551  | 
val concl = HOLogic.mk_Trueprop (mk_Bex field_suc_bd  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
552  | 
(Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs))));  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
553  | 
val vars = fold (Variable.add_free_names lthy) [prem, concl] [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
554  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
555  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies ([prem], concl))  | 
| 60728 | 556  | 
          (fn {context = ctxt, prems = _} => mk_bd_limit_tac ctxt n suc_bd_Cinfinite)
 | 
| 70494 | 557  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
558  | 
end;  | 
| 
 
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 timer = time (timer "Bounds");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
561  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
562  | 
(* minimal algebra *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
563  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
564  | 
fun mk_minG Asi i k = mk_UNION (mk_underS suc_bd $ i)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
565  | 
(Term.absfree jdx' (mk_nthN n (Asi $ jdx) k));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
566  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
567  | 
fun mk_minH_component Asi i sets Ts s k =  | 
| 69593 | 568  | 
HOLogic.mk_binop \<^const_name>\<open>sup\<close>  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
569  | 
(mk_minG Asi i k, mk_image s $ mk_in (passive_UNIVs @ map (mk_minG Asi i) ks) sets Ts);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
570  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
571  | 
fun mk_min_algs ss =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
572  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
573  | 
val BTs = map (range_type o fastype_of) ss;  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
574  | 
val Ts = passiveAs @ BTs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
575  | 
val (Asi, Asi') = `Free (Asi_name, suc_bdT -->  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
576  | 
Library.foldr1 HOLogic.mk_prodT (map HOLogic.mk_setT BTs));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
577  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
578  | 
mk_worec suc_bd (Term.absfree Asi' (Term.absfree idx' (HOLogic.mk_tuple  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
579  | 
           (@{map 4} (mk_minH_component Asi idx) (mk_setss Ts) (mk_FTs Ts) ss ks))))
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
580  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
581  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
582  | 
val (min_algs_thms, min_algs_mono_thms, card_of_min_algs_thm, least_min_algs_thm) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
583  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
584  | 
val i_field = HOLogic.mk_mem (idx, field_suc_bd);  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
585  | 
val min_algs = mk_min_algs ss;  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55770 
diff
changeset
 | 
586  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
587  | 
val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
588  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
589  | 
val concl = HOLogic.mk_Trueprop  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
590  | 
(HOLogic.mk_eq (min_algs $ idx, HOLogic.mk_tuple  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
591  | 
            (@{map 4} (mk_minH_component min_algs idx) setssAs FTsAs ss ks)));
 | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
592  | 
val goal = Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
593  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
594  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
595  | 
val min_algs_thm = Goal.prove_sorry lthy vars [] goal  | 
| 60728 | 596  | 
          (fn {context = ctxt, prems = _} => mk_min_algs_tac ctxt suc_bd_worel in_cong'_thms)
 | 
| 70494 | 597  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
598  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
599  | 
val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
600  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
601  | 
fun mk_mono_goal min_alg =  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
602  | 
HOLogic.mk_Trueprop (mk_relChain suc_bd (Term.absfree idx' min_alg));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
603  | 
|
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
604  | 
val monos =  | 
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
605  | 
map2 (fn goal => fn min_algs =>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
606  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
607  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
608  | 
              (fn {context = ctxt, prems = _} => mk_min_algs_mono_tac ctxt min_algs))
 | 
| 70494 | 609  | 
|> Thm.close_derivation \<^here>)  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
610  | 
(map mk_mono_goal min_algss) min_algs_thms;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
611  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
612  | 
fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
613  | 
val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss);  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
614  | 
val card_cT = Thm.ctyp_of lthy suc_bdT;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
615  | 
val card_ct = Thm.cterm_of lthy (Term.absfree idx' card_conjunction);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
616  | 
|
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
617  | 
val card_of =  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
618  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
619  | 
val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction));  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
620  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
621  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
622  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
623  | 
              (fn {context = ctxt, prems = _} => mk_min_algs_card_of_tac ctxt card_cT card_ct
 | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
624  | 
m suc_bd_worel min_algs_thms in_sbds  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
625  | 
sbd_Card_order sbd_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
626  | 
suc_bd_Asuc_bd Asuc_bd_Cinfinite)  | 
| 70494 | 627  | 
|> Thm.close_derivation \<^here>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
628  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
629  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
630  | 
val least_prem = HOLogic.mk_Trueprop (mk_alg Bs ss);  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51869 
diff
changeset
 | 
631  | 
val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs);  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
632  | 
val least_cT = Thm.ctyp_of lthy suc_bdT;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
633  | 
val least_ct = Thm.cterm_of lthy (Term.absfree idx' least_conjunction);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
634  | 
|
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
635  | 
val least =  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
636  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
637  | 
val goal = Logic.mk_implies (least_prem,  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
638  | 
HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction)));  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
639  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
640  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
641  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
642  | 
              (fn {context = ctxt, prems = _} => mk_min_algs_least_tac ctxt least_cT least_ct
 | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
643  | 
suc_bd_worel min_algs_thms alg_set_thms)  | 
| 70494 | 644  | 
|> Thm.close_derivation \<^here>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
645  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
646  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
647  | 
(min_algs_thms, monos, card_of, least)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
648  | 
end;  | 
| 
 
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  | 
val timer = time (timer "min_algs definition & thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
651  | 
|
| 53566 | 652  | 
val min_alg_binds = mk_internal_bs min_algN;  | 
653  | 
fun min_alg_bind i = nth min_alg_binds (i - 1);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
654  | 
val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
655  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
656  | 
fun min_alg_spec i =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
657  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
658  | 
val rhs = mk_UNION (field_suc_bd)  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
659  | 
(Term.absfree idx' (mk_nthN n (mk_min_algs ss $ idx) i));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
660  | 
in  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
661  | 
fold_rev (Term.absfree o Term.dest_Free) ss rhs  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
662  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
663  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
664  | 
val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) =  | 
| 59794 | 665  | 
lthy  | 
| 72450 | 666  | 
|> (snd o Local_Theory.begin_nested)  | 
| 59794 | 667  | 
|> fold_map (fn i => Local_Theory.define  | 
668  | 
((min_alg_bind i, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks  | 
|
669  | 
|>> apsnd split_list o split_list  | 
|
| 72450 | 670  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
671  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
672  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
673  | 
val min_algs = map (dest_Const_name o Morphism.term phi) min_alg_frees;  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
674  | 
val min_alg_defs = map (fn def =>  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
675  | 
mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) min_alg_def_frees;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
676  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
677  | 
fun mk_min_alg ss i =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
678  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
679  | 
val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1))))  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
680  | 
val Ts = map fastype_of ss;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
681  | 
val min_algT = Library.foldr (op -->) (Ts, T);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
682  | 
in  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
683  | 
Term.list_comb (Const (nth min_algs (i - 1), min_algT), ss)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
684  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
685  | 
|
| 56237 | 686  | 
val min_algs = map (mk_min_alg ss) ks;  | 
687  | 
||
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
688  | 
val ((Bs, ss), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
689  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
690  | 
|> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
691  | 
||>> mk_Frees "s" sTs;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
692  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
693  | 
val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
694  | 
let  | 
| 75624 | 695  | 
        val set_sbdss_weak = map (map (fn thm => @{thm ordLess_imp_ordLeq} OF [thm])) set_sbdss;
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
696  | 
val alg_min_alg =  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
697  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
698  | 
val goal = HOLogic.mk_Trueprop (mk_alg min_algs ss);  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
699  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
700  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
701  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
702  | 
              (fn {context = ctxt, prems = _} => mk_alg_min_alg_tac ctxt m alg_def min_alg_defs
 | 
| 75624 | 703  | 
suc_bd_limit_thm sbd_Cinfinite set_sbdss_weak min_algs_thms min_algs_mono_thms)  | 
| 70494 | 704  | 
|> Thm.close_derivation \<^here>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
705  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
706  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
707  | 
fun mk_card_of_thm min_alg def =  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
708  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
709  | 
val goal = HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd);  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
710  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
711  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
712  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
713  | 
              (fn {context = ctxt, prems = _} => mk_card_of_min_alg_tac ctxt def card_of_min_algs_thm
 | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
714  | 
suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite)  | 
| 70494 | 715  | 
|> Thm.close_derivation \<^here>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
716  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
717  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
718  | 
fun mk_least_thm min_alg B def =  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
719  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
720  | 
val prem = HOLogic.mk_Trueprop (mk_alg Bs ss);  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
721  | 
val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq min_alg B));  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
722  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
723  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
724  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
725  | 
              (fn {context = ctxt, prems = _} => mk_least_min_alg_tac ctxt def least_min_algs_thm)
 | 
| 70494 | 726  | 
|> Thm.close_derivation \<^here>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
727  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
728  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
729  | 
        val leasts = @{map 3} mk_least_thm min_algs Bs min_alg_defs;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
730  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
731  | 
val incl =  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
732  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
733  | 
val prem = HOLogic.mk_Trueprop (mk_alg Bs ss);  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
734  | 
val goal = Logic.mk_implies (prem,  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
735  | 
HOLogic.mk_Trueprop (mk_mor min_algs ss Bs ss active_ids));  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
736  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
737  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
738  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
739  | 
              (fn {context = ctxt, prems = _} =>
 | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
740  | 
EVERY' (rtac ctxt mor_incl_thm :: map (etac ctxt) leasts) 1)  | 
| 70494 | 741  | 
|> Thm.close_derivation \<^here>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
742  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
743  | 
in  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
744  | 
(alg_min_alg, map2 mk_card_of_thm min_algs min_alg_defs, leasts, incl)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
745  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
746  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
747  | 
val timer = time (timer "Minimal algebra definition & thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
748  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
749  | 
val II_repT = HOLogic.mk_prodT (HOLogic.mk_tupleT II_BTs, HOLogic.mk_tupleT II_sTs);  | 
| 53566 | 750  | 
val IIT_bind = mk_internal_b IITN;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
751  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
752  | 
val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) =  | 
| 49835 | 753  | 
typedef (IIT_bind, params, NoSyn)  | 
| 60728 | 754  | 
(HOLogic.mk_UNIV II_repT) NONE (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
 | 
755  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
756  | 
val IIT = Type (IIT_name, params');  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
757  | 
val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
758  | 
val Rep_IIT = Const (#Rep_name IIT_glob_info, IIT --> II_repT);  | 
| 49228 | 759  | 
val Abs_IIT_inverse_thm = UNIV_I RS #Abs_inverse IIT_loc_info;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
760  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
761  | 
val initT = IIT --> Asuc_bdT;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
762  | 
val active_initTs = replicate n initT;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
763  | 
val init_FTs = map2 (fn Ds => mk_T_of_bnf Ds (passiveAs @ active_initTs)) Dss bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
764  | 
val init_fTs = map (fn T => initT --> T) activeAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
765  | 
|
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
766  | 
val ((((II_Bs, II_ss), (iidx, iidx')), init_xFs), _) =  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
767  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
768  | 
|> mk_Frees "IIB" II_BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
769  | 
||>> mk_Frees "IIs" II_sTs  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
770  | 
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
771  | 
||>> mk_Frees "x" init_FTs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
772  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
773  | 
val II = HOLogic.mk_Collect (fst iidx', IIT, list_exists_free (II_Bs @ II_ss)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
774  | 
(HOLogic.mk_conj (HOLogic.mk_eq (iidx,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
775  | 
Abs_IIT $ (HOLogic.mk_prod (HOLogic.mk_tuple II_Bs, HOLogic.mk_tuple II_ss))),  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
776  | 
mk_alg II_Bs II_ss)));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
777  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
778  | 
val select_Bs = map (mk_nthN n (HOLogic.mk_fst (Rep_IIT $ iidx))) ks;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
779  | 
val select_ss = map (mk_nthN n (HOLogic.mk_snd (Rep_IIT $ iidx))) ks;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
780  | 
|
| 53566 | 781  | 
val str_init_binds = mk_internal_bs str_initN;  | 
782  | 
fun str_init_bind i = nth str_init_binds (i - 1);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
783  | 
val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
784  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
785  | 
fun str_init_spec i =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
786  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
787  | 
val init_xF = nth init_xFs (i - 1)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
788  | 
val select_s = nth select_ss (i - 1);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
789  | 
val map = mk_map_of_bnf (nth Dss (i - 1))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
790  | 
(passiveAs @ active_initTs) (passiveAs @ replicate n Asuc_bdT)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
791  | 
(nth bnfs (i - 1));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
792  | 
val map_args = passive_ids @ replicate n (mk_rapp iidx Asuc_bdT);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
793  | 
val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
794  | 
in  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
795  | 
fold_rev (Term.absfree o Term.dest_Free) [init_xF, iidx] rhs  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
796  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
797  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
798  | 
val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
799  | 
lthy  | 
| 72450 | 800  | 
|> (snd o Local_Theory.begin_nested)  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
801  | 
|> fold_map (fn i => Local_Theory.define  | 
| 
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
802  | 
((str_init_bind i, NoSyn), (str_init_def_bind i, str_init_spec i))) ks  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
803  | 
|>> apsnd split_list o split_list  | 
| 72450 | 804  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
805  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
806  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
807  | 
val str_inits =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
808  | 
map (Term.subst_atomic_types (map (`(Morphism.typ phi)) params') o Morphism.term phi)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
809  | 
str_init_frees;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
810  | 
|
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
811  | 
val str_init_defs = map (fn def =>  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
812  | 
mk_unabs_def 2 (HOLogic.mk_obj_eq (Morphism.thm phi def))) str_init_def_frees;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
813  | 
|
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
814  | 
val car_inits = map (mk_min_alg str_inits) ks;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
815  | 
|
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
816  | 
val (((((((((Bs, ss), Asuc_fs), (iidx, iidx')), init_xs), (init_xFs, init_xFs')), init_fs),  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
817  | 
init_fs_copy), init_phis), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
818  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
819  | 
|> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
820  | 
||>> mk_Frees "s" sTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
821  | 
||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs)  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
822  | 
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
823  | 
||>> mk_Frees "ix" active_initTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
824  | 
||>> mk_Frees' "x" init_FTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
825  | 
||>> mk_Frees "f" init_fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
826  | 
||>> mk_Frees "f" init_fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
827  | 
||>> mk_Frees "P" (replicate n (mk_pred1T initT));  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
828  | 
|
| 60784 | 829  | 
val alg_init_thm =  | 
830  | 
infer_instantiate' lthy (map (SOME o Thm.cterm_of lthy) str_inits) alg_min_alg_thm;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
831  | 
|
| 51551 | 832  | 
val alg_select_thm = Goal.prove_sorry lthy [] []  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
833  | 
(HOLogic.mk_Trueprop (mk_Ball II  | 
| 
55541
 
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
 
traytel 
parents: 
55538 
diff
changeset
 | 
834  | 
(Term.absfree iidx' (mk_alg select_Bs select_ss))))  | 
| 55197 | 835  | 
      (fn {context = ctxt, prems = _} => mk_alg_select_tac ctxt Abs_IIT_inverse_thm)
 | 
| 70494 | 836  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
837  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
838  | 
val mor_select_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
839  | 
let  | 
| 57567 | 840  | 
val i_prem = mk_Trueprop_mem (iidx, II);  | 
| 56237 | 841  | 
val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss active_UNIVs ss Asuc_fs);  | 
842  | 
val prems = [i_prem, mor_prem];  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
843  | 
val concl = HOLogic.mk_Trueprop  | 
| 56237 | 844  | 
(mk_mor car_inits str_inits active_UNIVs ss  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
845  | 
(map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs));  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
846  | 
val vars = fold (Variable.add_free_names lthy) (concl :: prems) [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
847  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
848  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl))  | 
| 60728 | 849  | 
          (fn {context = ctxt, prems = _} => mk_mor_select_tac ctxt mor_def mor_cong_thm
 | 
850  | 
mor_comp_thm mor_incl_min_alg_thm alg_def alg_select_thm alg_set_thms set_mapss  | 
|
851  | 
str_init_defs)  | 
|
| 70494 | 852  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
853  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
854  | 
|
| 56237 | 855  | 
val init_unique_mor_thms =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
856  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
857  | 
val prems = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_mem) init_xs car_inits  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
858  | 
val mor_prems = map HOLogic.mk_Trueprop  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
859  | 
[mk_mor car_inits str_inits Bs ss init_fs,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
860  | 
mk_mor car_inits str_inits Bs ss init_fs_copy];  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
861  | 
fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
862  | 
val unique = HOLogic.mk_Trueprop  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
863  | 
          (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_fun_eq init_fs init_fs_copy init_xs));
 | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
864  | 
val cts = map (Thm.cterm_of lthy) ss;  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
865  | 
val all_prems = prems @ mor_prems;  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
866  | 
val vars = fold (Variable.add_free_names lthy) (unique :: all_prems) [];  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
867  | 
val unique_mor =  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
868  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (all_prems, unique))  | 
| 60728 | 869  | 
            (fn {context = ctxt, prems = _} => mk_init_unique_mor_tac ctxt cts m alg_def
 | 
870  | 
alg_init_thm least_min_alg_thms in_mono'_thms alg_set_thms morE_thms map_cong0s)  | 
|
| 70494 | 871  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
872  | 
in  | 
| 56237 | 873  | 
split_conj_thm unique_mor  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
874  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
875  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
876  | 
val init_setss = mk_setss (passiveAs @ active_initTs);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
877  | 
val active_init_setss = map (drop m) init_setss;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
878  | 
val init_ins = map2 (fn sets => mk_in (passive_UNIVs @ car_inits) sets) init_setss init_FTs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
879  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
880  | 
fun mk_closed phis =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
881  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
882  | 
fun mk_conjunct phi str_init init_sets init_in x x' =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
883  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
884  | 
val prem = Library.foldr1 HOLogic.mk_conj  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
885  | 
(map2 (fn set => mk_Ball (set $ x)) init_sets phis);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
886  | 
val concl = phi $ (str_init $ x);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
887  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
888  | 
mk_Ball init_in (Term.absfree x' (HOLogic.mk_imp (prem, concl)))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
889  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
890  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
891  | 
Library.foldr1 HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
892  | 
          (@{map 6} mk_conjunct phis str_inits active_init_setss init_ins init_xFs init_xFs')
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
893  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
894  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
895  | 
val init_induct_thm =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
896  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
897  | 
val prem = HOLogic.mk_Trueprop (mk_closed init_phis);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
898  | 
val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
899  | 
(map2 mk_Ball car_inits init_phis));  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
900  | 
val vars = fold (Variable.add_free_names lthy) [concl, prem] [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
901  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
902  | 
Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl))  | 
| 60728 | 903  | 
          (fn {context = ctxt, prems = _} => mk_init_induct_tac ctxt m alg_def alg_init_thm
 | 
904  | 
least_min_alg_thms alg_set_thms)  | 
|
| 70494 | 905  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
906  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
907  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
908  | 
val timer = time (timer "Initiality definition & thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
909  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
910  | 
val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
911  | 
lthy  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
912  | 
      |> @{fold_map 3} (fn b => fn mx => fn car_init =>
 | 
| 57093 | 913  | 
typedef (b, params, mx) car_init NONE  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
914  | 
(fn ctxt =>  | 
| 60728 | 915  | 
            EVERY' [rtac ctxt iffD2, rtac ctxt @{thm ex_in_conv}, resolve_tac ctxt alg_not_empty_thms,
 | 
916  | 
rtac ctxt alg_init_thm] 1)) bs mixfixes car_inits  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
917  | 
|>> apsnd split_list o split_list;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
918  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
919  | 
val Ts = map (fn name => Type (name, params')) T_names;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
920  | 
fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
921  | 
val Ts' = mk_Ts passiveBs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
922  | 
val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> initT)) T_glob_infos Ts;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
923  | 
val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, initT --> T)) T_glob_infos Ts;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
924  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
925  | 
val type_defs = map #type_definition T_loc_infos;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
926  | 
val Reps = map #Rep T_loc_infos;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
927  | 
val Rep_inverses = map #Rep_inverse T_loc_infos;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
928  | 
val Abs_inverses = map #Abs_inverse T_loc_infos;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
929  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
930  | 
val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
931  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
932  | 
val UNIVs = map HOLogic.mk_UNIV Ts;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
933  | 
val FTs = mk_FTs (passiveAs @ Ts);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
934  | 
val FTs' = mk_FTs (passiveBs @ Ts');  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
935  | 
fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
936  | 
val setFTss = map (mk_FTs o mk_set_Ts) passiveAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
937  | 
val FTs_setss = mk_setss (passiveAs @ Ts);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
938  | 
val FTs'_setss = mk_setss (passiveBs @ Ts');  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
939  | 
val map_FT_inits = map2 (fn Ds =>  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
940  | 
mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ active_initTs)) Dss bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
941  | 
val fTs = map2 (curry op -->) Ts activeAs;  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
942  | 
val foldT = Library.foldr1 HOLogic.mk_prodT (map2 (curry op -->) Ts activeAs);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
943  | 
|
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
944  | 
val ((ss, (fold_f, fold_f')), _) =  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
945  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
946  | 
|> mk_Frees "s" sTs  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
947  | 
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT;  | 
| 
49330
 
276ff43ee0b1
reuse generated names (they look better + slightly more efficient)
 
blanchet 
parents: 
49327 
diff
changeset
 | 
948  | 
|
| 54492 | 949  | 
fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");  | 
| 59859 | 950  | 
val ctor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o ctor_bind;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
951  | 
|
| 55756 | 952  | 
fun ctor_spec abs str map_FT_init =  | 
953  | 
Library.foldl1 HOLogic.mk_comp [abs, str,  | 
|
954  | 
Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts)];  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
955  | 
|
| 49501 | 956  | 
val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =  | 
| 49311 | 957  | 
lthy  | 
| 72450 | 958  | 
|> (snd o Local_Theory.begin_nested)  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
959  | 
      |> @{fold_map 4} (fn i => fn abs => fn str => fn mapx =>
 | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
960  | 
Local_Theory.define  | 
| 55756 | 961  | 
((ctor_bind i, NoSyn), (ctor_def_bind i, ctor_spec abs str mapx)))  | 
962  | 
ks Abs_Ts str_inits map_FT_inits  | 
|
| 49311 | 963  | 
|>> apsnd split_list o split_list  | 
| 72450 | 964  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
965  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
966  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 49501 | 967  | 
fun mk_ctors passive =  | 
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49176 
diff
changeset
 | 
968  | 
map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o  | 
| 49501 | 969  | 
Morphism.term phi) ctor_frees;  | 
970  | 
val ctors = mk_ctors passiveAs;  | 
|
971  | 
val ctor's = mk_ctors passiveBs;  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
972  | 
val ctor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) ctor_def_frees;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
973  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
974  | 
val (mor_Rep_thm, mor_Abs_thm) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
975  | 
let  | 
| 56237 | 976  | 
val defs = mor_def :: ctor_defs;  | 
977  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
978  | 
val mor_Rep =  | 
| 51551 | 979  | 
Goal.prove_sorry lthy [] []  | 
| 49501 | 980  | 
(HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts))  | 
| 56237 | 981  | 
            (fn {context = ctxt, prems = _} => mk_mor_Rep_tac ctxt m defs Reps Abs_inverses
 | 
982  | 
alg_min_alg_thm alg_set_thms set_mapss)  | 
|
| 70494 | 983  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
984  | 
|
| 56237 | 985  | 
fun mk_ct initFT str abs = Term.absdummy initFT (abs $ (str $ Bound 0))  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
986  | 
        val cts = @{map 3} (Thm.cterm_of lthy ooo mk_ct) init_FTs str_inits Abs_Ts;
 | 
| 56237 | 987  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
988  | 
val mor_Abs =  | 
| 51551 | 989  | 
Goal.prove_sorry lthy [] []  | 
| 49501 | 990  | 
(HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts))  | 
| 56237 | 991  | 
            (fn {context = ctxt, prems = _} => mk_mor_Abs_tac ctxt cts defs Abs_inverses
 | 
992  | 
map_comp_id_thms map_cong0L_thms)  | 
|
| 70494 | 993  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
994  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
995  | 
(mor_Rep, mor_Abs)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
996  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
997  | 
|
| 49501 | 998  | 
val timer = time (timer "ctor definitions & thms");  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
999  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1000  | 
val fold_fun = Term.absfree fold_f'  | 
| 
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1001  | 
(mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks));  | 
| 
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1002  | 
val foldx = HOLogic.choice_const foldT $ fold_fun;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1003  | 
|
| 54492 | 1004  | 
fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_");  | 
| 59859 | 1005  | 
val fold_def_bind = rpair [] o Binding.concealed o Thm.def_binding o fold_bind;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1006  | 
|
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1007  | 
fun fold_spec i = fold_rev (Term.absfree o Term.dest_Free) ss (mk_nthN n foldx i);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1008  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1009  | 
val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) =  | 
| 49311 | 1010  | 
lthy  | 
| 72450 | 1011  | 
|> (snd o Local_Theory.begin_nested)  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1012  | 
|> fold_map (fn i =>  | 
| 
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1013  | 
Local_Theory.define ((fold_bind i, NoSyn), (fold_def_bind i, fold_spec i))) ks  | 
| 49311 | 1014  | 
|>> apsnd split_list o split_list  | 
| 72450 | 1015  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1016  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1017  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1018  | 
val folds = map (Morphism.term phi) fold_frees;  | 
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
1019  | 
val fold_names = map dest_Const_name folds;  | 
| 52731 | 1020  | 
fun mk_folds passives actives =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1021  | 
      @{map 3} (fn name => fn T => fn active =>
 | 
| 52731 | 1022  | 
Const (name, Library.foldr (op -->)  | 
| 52923 | 1023  | 
(map2 (curry op -->) (mk_FTs (passives @ actives)) actives, T --> active)))  | 
| 52731 | 1024  | 
fold_names (mk_Ts passives) actives;  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1025  | 
fun mk_fold Ts ss i = Term.list_comb (Const (nth fold_names (i - 1), Library.foldr (op -->)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1026  | 
(map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1027  | 
val fold_defs = map (fn def =>  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
1028  | 
mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) fold_def_frees;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1029  | 
|
| 56237 | 1030  | 
(* algebra copies *)  | 
1031  | 
||
| 63312 | 1032  | 
val ((((((Bs, B's), ss), s's), inv_fs), fs), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1033  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1034  | 
|> mk_Frees "B" BTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1035  | 
||>> mk_Frees "B'" B'Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1036  | 
||>> mk_Frees "s" sTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1037  | 
||>> mk_Frees "s'" s'Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1038  | 
||>> mk_Frees "f" inv_fTs  | 
| 63312 | 1039  | 
||>> mk_Frees "f" fTs;  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1040  | 
|
| 56237 | 1041  | 
val copy_thm =  | 
1042  | 
let  | 
|
1043  | 
val prems = HOLogic.mk_Trueprop (mk_alg Bs ss) ::  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1044  | 
          @{map 3} (HOLogic.mk_Trueprop ooo mk_bij_betw) inv_fs B's Bs;
 | 
| 56237 | 1045  | 
val concl = HOLogic.mk_Trueprop (list_exists_free s's  | 
1046  | 
(HOLogic.mk_conj (mk_alg B's s's, mk_mor B's s's Bs ss inv_fs)));  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1047  | 
val vars = fold (Variable.add_free_names lthy) (concl :: prems) [];  | 
| 56237 | 1048  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1049  | 
Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl))  | 
| 60728 | 1050  | 
          (fn {context = ctxt, prems = _} => mk_copy_tac ctxt m alg_def mor_def alg_set_thms
 | 
1051  | 
set_mapss)  | 
|
| 70494 | 1052  | 
|> Thm.close_derivation \<^here>  | 
| 56237 | 1053  | 
end;  | 
1054  | 
||
1055  | 
val init_ex_mor_thm =  | 
|
1056  | 
let  | 
|
1057  | 
val goal = HOLogic.mk_Trueprop  | 
|
1058  | 
(list_exists_free fs (mk_mor UNIVs ctors active_UNIVs ss fs));  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1059  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 56237 | 1060  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1061  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1062  | 
          (fn {context = ctxt, prems = _} =>
 | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1063  | 
mk_init_ex_mor_tac ctxt Abs_IIT_inverse_thm (alg_min_alg_thm RS copy_thm)  | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1064  | 
card_of_min_alg_thms mor_Rep_thm mor_comp_thm mor_select_thm mor_incl_thm)  | 
| 70494 | 1065  | 
|> Thm.close_derivation \<^here>  | 
| 56237 | 1066  | 
end;  | 
1067  | 
||
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1068  | 
val mor_fold_thm =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1069  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1070  | 
val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks);  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1071  | 
val cT = Thm.ctyp_of lthy foldT;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1072  | 
val ct = Thm.cterm_of lthy fold_fun  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1073  | 
val goal = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks));  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1074  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1075  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1076  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
1077  | 
          (fn {context = ctxt, ...} =>
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
1078  | 
mk_mor_fold_tac ctxt cT ct fold_defs init_ex_mor_thm mor_cong)  | 
| 70494 | 1079  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1080  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1081  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1082  | 
val ctor_fold_thms = map (fn morE => rule_by_tactic lthy  | 
| 60728 | 1083  | 
      ((rtac lthy CollectI THEN' CONJ_WRAP' (K (rtac lthy @{thm subset_UNIV})) (1 upto m + n)) 1)
 | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1084  | 
(mor_fold_thm RS morE)) morE_thms;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1085  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1086  | 
val (fold_unique_mor_thms, fold_unique_mor_thm) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1087  | 
let  | 
| 49501 | 1088  | 
val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss fs);  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1089  | 
fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1090  | 
val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1091  | 
val vars = fold (Variable.add_free_names lthy) [prem, unique] [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1092  | 
val unique_mor = Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, unique))  | 
| 60728 | 1093  | 
          (fn {context = ctxt, prems = _} => mk_fold_unique_mor_tac ctxt type_defs
 | 
1094  | 
init_unique_mor_thms Reps mor_comp_thm mor_Abs_thm mor_fold_thm)  | 
|
| 70494 | 1095  | 
|> Thm.close_derivation \<^here>;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1096  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1097  | 
`split_conj_thm unique_mor  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1098  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1099  | 
|
| 52911 | 1100  | 
val (ctor_fold_unique_thms, ctor_fold_unique_thm) =  | 
1101  | 
`split_conj_thm (mk_conjIN n RS  | 
|
| 52904 | 1102  | 
(mor_UNIV_thm RS iffD2 RS fold_unique_mor_thm))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1103  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1104  | 
val fold_ctor_thms =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1105  | 
      map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym)
 | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1106  | 
fold_unique_mor_thms;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1107  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1108  | 
val ctor_o_fold_thms =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1109  | 
let  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1110  | 
val mor = mor_comp_thm OF [mor_fold_thm, mor_str_thm];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1111  | 
in  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1112  | 
map2 (fn unique => fn fold_ctor =>  | 
| 
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1113  | 
trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1114  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1115  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1116  | 
val timer = time (timer "fold definitions & thms");  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1117  | 
|
| 49501 | 1118  | 
val map_ctors = map2 (fn Ds => fn bnf =>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1119  | 
Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf,  | 
| 49501 | 1120  | 
map HOLogic.id_const passiveAs @ ctors)) Dss bnfs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1121  | 
|
| 54492 | 1122  | 
fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");  | 
| 59859 | 1123  | 
val dtor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o dtor_bind;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1124  | 
|
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1125  | 
fun dtor_spec i = mk_fold Ts map_ctors i;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1126  | 
|
| 49501 | 1127  | 
val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =  | 
| 49311 | 1128  | 
lthy  | 
| 72450 | 1129  | 
|> (snd o Local_Theory.begin_nested)  | 
| 
55204
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1130  | 
|> fold_map (fn i =>  | 
| 
 
345ee77213b5
use Local_Theory.define instead of Specification.definition for internal constants
 
traytel 
parents: 
55197 
diff
changeset
 | 
1131  | 
Local_Theory.define ((dtor_bind i, NoSyn), (dtor_def_bind i, dtor_spec i))) ks  | 
| 49311 | 1132  | 
|>> apsnd split_list o split_list  | 
| 72450 | 1133  | 
||> `Local_Theory.end_nested;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1134  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1135  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 49501 | 1136  | 
fun mk_dtors params =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1137  | 
map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)  | 
| 49501 | 1138  | 
dtor_frees;  | 
1139  | 
val dtors = mk_dtors params';  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
1140  | 
val dtor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) dtor_def_frees;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1141  | 
|
| 
63222
 
fe92356ade42
eliminated pointless alias (no warning for duplicates);
 
wenzelm 
parents: 
63045 
diff
changeset
 | 
1142  | 
val ctor_o_dtor_thms = map2 (Local_Defs.fold lthy o single) dtor_defs ctor_o_fold_thms;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1143  | 
|
| 49501 | 1144  | 
val dtor_o_ctor_thms =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1145  | 
let  | 
| 49501 | 1146  | 
fun mk_goal dtor ctor FT =  | 
1147  | 
mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1148  | 
        val goals = @{map 3} mk_goal dtors ctors FTs;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1149  | 
in  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1150  | 
        @{map 5} (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L =>
 | 
| 51551 | 1151  | 
Goal.prove_sorry lthy [] [] goal  | 
| 60728 | 1152  | 
            (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt dtor_def foldx map_comp_id
 | 
1153  | 
map_cong0L ctor_o_fold_thms)  | 
|
| 70494 | 1154  | 
|> Thm.close_derivation \<^here>)  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
1155  | 
goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1156  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1157  | 
|
| 49501 | 1158  | 
    val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
 | 
1159  | 
    val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1160  | 
|
| 49501 | 1161  | 
val bij_dtor_thms =  | 
1162  | 
      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
 | 
|
1163  | 
    val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
 | 
|
1164  | 
    val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
 | 
|
1165  | 
    val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
 | 
|
1166  | 
    val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
 | 
|
1167  | 
val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1168  | 
|
| 49501 | 1169  | 
val bij_ctor_thms =  | 
1170  | 
      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
 | 
|
1171  | 
    val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
 | 
|
1172  | 
    val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
 | 
|
1173  | 
    val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
 | 
|
1174  | 
    val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
 | 
|
1175  | 
val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1176  | 
|
| 49501 | 1177  | 
val timer = time (timer "dtor definitions & thms");  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1178  | 
|
| 63312 | 1179  | 
val (((((((Izs, (Izs1, Izs1'))), (Izs2, Izs2')), xFs), yFs), init_phis), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1180  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1181  | 
|> mk_Frees "z" Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1182  | 
||>> mk_Frees' "z1" Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1183  | 
||>> mk_Frees' "z2" Ts'  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1184  | 
||>> mk_Frees "x" FTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1185  | 
||>> mk_Frees "y" FTs'  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1186  | 
||>> mk_Frees "P" (replicate n (mk_pred1T initT));  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1187  | 
|
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1188  | 
val phis = map2 retype_const_or_free (map mk_pred1T Ts) init_phis;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1189  | 
val phi2s = map2 retype_const_or_free (map2 mk_pred2T Ts Ts') init_phis;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1190  | 
|
| 49501 | 1191  | 
val (ctor_induct_thm, induct_params) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1192  | 
let  | 
| 49501 | 1193  | 
fun mk_prem phi ctor sets x =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1194  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1195  | 
fun mk_IH phi set z =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1196  | 
let  | 
| 57567 | 1197  | 
val prem = mk_Trueprop_mem (z, set $ x);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1198  | 
val concl = HOLogic.mk_Trueprop (phi $ z);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1199  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1200  | 
Logic.all z (Logic.mk_implies (prem, concl))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1201  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1202  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1203  | 
            val IHs = @{map 3} mk_IH phis (drop m sets) Izs;
 | 
| 49501 | 1204  | 
val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1205  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1206  | 
Logic.all x (Logic.list_implies (IHs, concl))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1207  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1208  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1209  | 
        val prems = @{map 4} mk_prem phis ctors FTs_setss xFs;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1210  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1211  | 
fun mk_concl phi z = phi $ z;  | 
| 57307 | 1212  | 
val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_concl phis Izs));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1213  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1214  | 
val goal = Logic.list_implies (prems, concl);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1215  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1216  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1217  | 
(Goal.prove_sorry lthy vars [] goal  | 
| 62698 | 1218  | 
          (fn {context = ctxt, prems = _} =>
 | 
| 61271 | 1219  | 
mk_ctor_induct_tac ctxt m set_mapss init_induct_thm morE_thms mor_Abs_thm  | 
1220  | 
Rep_inverses Abs_inverses Reps)  | 
|
| 70494 | 1221  | 
|> Thm.close_derivation \<^here>,  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1222  | 
rev (Term.add_tfrees goal []))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1223  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1224  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1225  | 
val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct_params;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1226  | 
|
| 49501 | 1227  | 
val weak_ctor_induct_thms =  | 
| 54487 | 1228  | 
let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI);  | 
| 49501 | 1229  | 
in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1230  | 
|
| 49501 | 1231  | 
val (ctor_induct2_thm, induct2_params) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1232  | 
let  | 
| 49501 | 1233  | 
fun mk_prem phi ctor ctor' sets sets' x y =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1234  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1235  | 
fun mk_IH phi set set' z1 z2 =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1236  | 
let  | 
| 57567 | 1237  | 
val prem1 = mk_Trueprop_mem (z1, (set $ x));  | 
1238  | 
val prem2 = mk_Trueprop_mem (z2, (set' $ y));  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1239  | 
val concl = HOLogic.mk_Trueprop (phi $ z1 $ z2);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1240  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1241  | 
fold_rev Logic.all [z1, z2] (Logic.list_implies ([prem1, prem2], concl))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1242  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1243  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1244  | 
            val IHs = @{map 5} mk_IH phi2s (drop m sets) (drop m sets') Izs1 Izs2;
 | 
| 49501 | 1245  | 
val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x) $ (ctor' $ y));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1246  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1247  | 
fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1248  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1249  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1250  | 
        val prems = @{map 7} mk_prem phi2s ctors ctor's FTs_setss FTs'_setss xFs yFs;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1251  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1252  | 
fun mk_concl phi z1 z2 = phi $ z1 $ z2;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1253  | 
val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1254  | 
          (@{map 3} mk_concl phi2s Izs1 Izs2));
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1255  | 
fun mk_t phi (z1, z1') (z2, z2') =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1256  | 
Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2));  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1257  | 
        val cts = @{map 3} (SOME o Thm.cterm_of lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2');
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1258  | 
val goal = Logic.list_implies (prems, concl);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1259  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1260  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1261  | 
(Goal.prove_sorry lthy vars [] goal  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1262  | 
          (fn {context = ctxt, prems = _} => mk_ctor_induct2_tac ctxt cTs cts ctor_induct_thm
 | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1263  | 
weak_ctor_induct_thms)  | 
| 70494 | 1264  | 
|> Thm.close_derivation \<^here>,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1265  | 
rev (Term.add_tfrees goal []))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1266  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1267  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1268  | 
val timer = time (timer "induction");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1269  | 
|
| 53270 | 1270  | 
fun mk_ctor_map_DEADID_thm ctor_inject map_id0 =  | 
1271  | 
trans OF [id_apply, iffD2 OF [ctor_inject, map_id0 RS sym]];  | 
|
| 
51917
 
f964a9887713
store proper theorems even for fixed points that have no passive live variables
 
traytel 
parents: 
51894 
diff
changeset
 | 
1272  | 
|
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1273  | 
fun mk_ctor_map_unique_DEADID_thm () =  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1274  | 
let  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1275  | 
val (funs, algs) =  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1276  | 
HOLogic.conjuncts (HOLogic.dest_Trueprop (Thm.concl_of ctor_fold_unique_thm))  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1277  | 
|> map_split HOLogic.dest_eq  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1278  | 
||> snd o strip_comb o hd  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1279  | 
          |> @{apply 2} (map (fst o dest_Var));
 | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1280  | 
fun mk_fun_insts T ix = Thm.cterm_of lthy (Var (ix, T --> T));  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1281  | 
val theta =  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1282  | 
          (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) ctors);
 | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1283  | 
val ctor_fold_ctors = (ctor_fold_unique_thm OF  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1284  | 
          map (fn thm => mk_trans @{thm id_o} (mk_sym (thm RS
 | 
| 67399 | 1285  | 
            @{thm trans[OF arg_cong2[of _ _ _ _ "(\<circ>)", OF refl] o_id]}))) map_id0s)
 | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1286  | 
|> split_conj_thm |> map mk_sym;  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1287  | 
in  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1288  | 
infer_instantiate lthy theta ctor_fold_unique_thm  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1289  | 
|> unfold_thms lthy ctor_fold_ctors  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1290  | 
|> Morphism.thm (Local_Theory.target_morphism lthy)  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1291  | 
end;  | 
| 
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1292  | 
|
| 
51917
 
f964a9887713
store proper theorems even for fixed points that have no passive live variables
 
traytel 
parents: 
51894 
diff
changeset
 | 
1293  | 
fun mk_ctor_Irel_DEADID_thm ctor_inject bnf =  | 
| 
 
f964a9887713
store proper theorems even for fixed points that have no passive live variables
 
traytel 
parents: 
51894 
diff
changeset
 | 
1294  | 
      trans OF [ctor_inject, rel_eq_of_bnf bnf RS @{thm predicate2_eqD} RS sym];
 | 
| 
 
f964a9887713
store proper theorems even for fixed points that have no passive live variables
 
traytel 
parents: 
51894 
diff
changeset
 | 
1295  | 
|
| 51918 | 1296  | 
val IphiTs = map2 mk_pred2T passiveAs passiveBs;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1297  | 
val Ipsi1Ts = map2 mk_pred2T passiveAs passiveCs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1298  | 
val Ipsi2Ts = map2 mk_pred2T passiveCs passiveBs;  | 
| 52731 | 1299  | 
val activephiTs = map2 mk_pred2T activeAs activeBs;  | 
| 51918 | 1300  | 
val activeIphiTs = map2 mk_pred2T Ts Ts';  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1301  | 
|
| 51918 | 1302  | 
val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;  | 
1303  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1304  | 
(*register new datatypes as BNFs*)  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1305  | 
val (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Imap_unique_thm, ctor_Iset_thmss',  | 
| 53567 | 1306  | 
ctor_Irel_thms, Ibnf_notes, lthy) =  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
1307  | 
if m = 0 then  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52911 
diff
changeset
 | 
1308  | 
(timer, replicate n DEADID_bnf,  | 
| 64413 | 1309  | 
map_split (`(mk_pointfree2 lthy)) (map2 mk_ctor_map_DEADID_thm ctor_inject_thms map_ids),  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1310  | 
mk_ctor_map_unique_DEADID_thm (),  | 
| 53567 | 1311  | 
replicate n [], map2 mk_ctor_Irel_DEADID_thm ctor_inject_thms bnfs, [], lthy)  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
1312  | 
else let  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1313  | 
val fTs = map2 (curry op -->) passiveAs passiveBs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1314  | 
val uTs = map2 (curry op -->) Ts Ts';  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1315  | 
|
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1316  | 
val ((((fs, fs'), (AFss, AFss')), (ys, ys')), _) =  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1317  | 
lthy  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1318  | 
|> mk_Frees' "f" fTs  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1319  | 
||>> mk_Freess' "z" setFTss  | 
| 51918 | 1320  | 
||>> mk_Frees' "y" passiveAs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1321  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1322  | 
val map_FTFT's = map2 (fn Ds =>  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1323  | 
mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1324  | 
fun mk_passive_maps ATs BTs Ts =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1325  | 
map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ Ts)) Dss bnfs;  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1326  | 
fun mk_map_fold_arg fs Ts ctor fmap =  | 
| 49501 | 1327  | 
HOLogic.mk_comp (ctor, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts));  | 
1328  | 
fun mk_map Ts fs Ts' ctors mk_maps =  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1329  | 
mk_fold Ts (map2 (mk_map_fold_arg fs Ts') ctors (mk_maps Ts'));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1330  | 
val pmapsABT' = mk_passive_maps passiveAs passiveBs;  | 
| 49501 | 1331  | 
val fs_maps = map (mk_map Ts fs Ts' ctor's pmapsABT') ks;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1332  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1333  | 
val ls = 1 upto m;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1334  | 
val setsss = map (mk_setss o mk_set_Ts) passiveAs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1335  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1336  | 
fun mk_col l T z z' sets =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1337  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1338  | 
fun mk_UN set = mk_Union T $ (set $ z);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1339  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1340  | 
Term.absfree z'  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1341  | 
(mk_union (nth sets (l - 1) $ z,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1342  | 
Library.foldl1 mk_union (map mk_UN (drop m sets))))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1343  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1344  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1345  | 
        val colss = @{map 5} (fn l => fn T => @{map 3} (mk_col l T)) ls passiveAs AFss AFss' setsss;
 | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
1346  | 
val setss_by_range = map (fn cols => map (mk_fold Ts cols) ks) colss;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1347  | 
val setss_by_bnf = transpose setss_by_range;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1348  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1349  | 
val set_bss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1350  | 
map (flat o map2 (fn B => fn b =>  | 
| 
55702
 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 
blanchet 
parents: 
55541 
diff
changeset
 | 
1351  | 
if member (op =) deads (TFree B) then [] else [b]) resBs) set_bss0;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1352  | 
|
| 49501 | 1353  | 
val ctor_witss =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1354  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1355  | 
val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1356  | 
(replicate (nwits_of_bnf bnf) Ds)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1357  | 
(replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1358  | 
fun close_wit (I, wit) = fold_rev Term.absfree (map (nth ys') I) wit;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1359  | 
fun wit_apply (arg_I, arg_wit) (fun_I, fun_wit) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1360  | 
(union (op =) arg_I fun_I, fun_wit $ arg_wit);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1361  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1362  | 
fun gen_arg support i =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1363  | 
if i < m then [([i], nth ys i)]  | 
| 49501 | 1364  | 
else maps (mk_wit support (nth ctors (i - m)) (i - m)) (nth support (i - m))  | 
1365  | 
and mk_wit support ctor i (I, wit) =  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1366  | 
let val args = map (gen_arg (nth_map i (remove (op =) (I, wit)) support)) I;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1367  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1368  | 
(args, [([], wit)])  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1369  | 
|-> fold (map_product wit_apply)  | 
| 49501 | 1370  | 
|> map (apsnd (fn t => ctor $ t))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1371  | 
|> minimize_wits  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1372  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1373  | 
in  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1374  | 
            @{map 3} (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i))
 | 
| 49501 | 1375  | 
ctors (0 upto n - 1) witss  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1376  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1377  | 
|
| 75624 | 1378  | 
val (lthy, sbd0, sbd0_card_order, sbd0_Cinfinite, sbd0_regularCard, set_sbd0ss) =  | 
| 56350 | 1379  | 
if n = 1  | 
| 75624 | 1380  | 
then (lthy, hd bd0s, hd bd0_card_orders, hd bd0_Cinfinites, hd bd0_regularCards, set_bd0ss)  | 
| 56350 | 1381  | 
else  | 
1382  | 
let  | 
|
1383  | 
val sum_bd0 = Library.foldr1 (uncurry mk_csum) bd0s;  | 
|
1384  | 
val sum_bd0T = fst (dest_relT (fastype_of sum_bd0));  | 
|
| 
56516
 
a13c2ccc160b
more accurate type arguments for intermeadiate typedefs
 
traytel 
parents: 
56350 
diff
changeset
 | 
1385  | 
val (sum_bd0T_params, sum_bd0T_params') = `(map TFree) (Term.add_tfreesT sum_bd0T []);  | 
| 56350 | 1386  | 
|
1387  | 
val sbd0T_bind = mk_internal_b (sum_bdTN ^ "0");  | 
|
| 57567 | 1388  | 
|
| 56350 | 1389  | 
val ((sbd0T_name, (sbd0T_glob_info, sbd0T_loc_info)), lthy) =  | 
| 
56516
 
a13c2ccc160b
more accurate type arguments for intermeadiate typedefs
 
traytel 
parents: 
56350 
diff
changeset
 | 
1390  | 
typedef (sbd0T_bind, sum_bd0T_params', NoSyn)  | 
| 60728 | 1391  | 
(HOLogic.mk_UNIV sum_bd0T) NONE (fn ctxt =>  | 
1392  | 
EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy;  | 
|
| 57567 | 1393  | 
|
| 
56516
 
a13c2ccc160b
more accurate type arguments for intermeadiate typedefs
 
traytel 
parents: 
56350 
diff
changeset
 | 
1394  | 
val sbd0T = Type (sbd0T_name, sum_bd0T_params);  | 
| 56350 | 1395  | 
val Abs_sbd0T = Const (#Abs_name sbd0T_glob_info, sum_bd0T --> sbd0T);  | 
| 57567 | 1396  | 
|
| 56350 | 1397  | 
val sbd0_bind = mk_internal_b (sum_bdN ^ "0");  | 
1398  | 
val sbd0_def_bind = (Thm.def_binding sbd0_bind, []);  | 
|
| 57567 | 1399  | 
|
| 56350 | 1400  | 
val sbd0_spec = mk_dir_image sum_bd0 Abs_sbd0T;  | 
| 57567 | 1401  | 
|
| 56350 | 1402  | 
val ((sbd0_free, (_, sbd0_def_free)), (lthy, lthy_old)) =  | 
1403  | 
lthy  | 
|
| 72450 | 1404  | 
|> (snd o Local_Theory.begin_nested)  | 
| 56350 | 1405  | 
|> Local_Theory.define ((sbd0_bind, NoSyn), (sbd0_def_bind, sbd0_spec))  | 
| 72450 | 1406  | 
||> `Local_Theory.end_nested;  | 
| 57567 | 1407  | 
|
| 56350 | 1408  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 57567 | 1409  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
1410  | 
val sbd0_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd0_def_free);  | 
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
1411  | 
val sbd0 = Const (dest_Const_name (Morphism.term phi sbd0_free), mk_relT (`I sbd0T));  | 
| 57567 | 1412  | 
|
| 56350 | 1413  | 
val Abs_sbd0T_inj = mk_Abs_inj_thm (#Abs_inject sbd0T_loc_info);  | 
1414  | 
val Abs_sbd0T_bij = mk_Abs_bij_thm lthy Abs_sbd0T_inj (#Abs_cases sbd0T_loc_info);  | 
|
| 57567 | 1415  | 
|
| 75624 | 1416  | 
val (sum_Cinfinite, sum_regularCard) =  | 
1417  | 
mk_sum_Cinfinite_regularCard (bd0_Cinfinites ~~ bd0_regularCards);  | 
|
| 56350 | 1418  | 
val sum_Card_order = sum_Cinfinite RS conjunct2;  | 
1419  | 
val sum_card_order = mk_sum_card_order bd0_card_orders;  | 
|
| 57567 | 1420  | 
|
| 56350 | 1421  | 
              val sbd0_ordIso = @{thm ssubst_Pair_rhs} OF
 | 
1422  | 
                [@{thm dir_image} OF [Abs_sbd0T_inj, sum_Card_order], sbd0_def];
 | 
|
1423  | 
              val sbd0_Cinfinite = @{thm Cinfinite_cong} OF [sbd0_ordIso, sum_Cinfinite];
 | 
|
| 57567 | 1424  | 
|
| 56350 | 1425  | 
              val sbd0_card_order = @{thm iffD2[OF arg_cong[of _ _ card_order]]} OF
 | 
1426  | 
                [sbd0_def, @{thm card_order_dir_image} OF [Abs_sbd0T_bij, sum_card_order]];
 | 
|
| 57567 | 1427  | 
|
| 75624 | 1428  | 
              val sbd0_regularCard = @{thm regularCard_ordIso} OF
 | 
1429  | 
[sbd0_ordIso, sum_Cinfinite, sum_regularCard];  | 
|
1430  | 
||
| 56350 | 1431  | 
fun mk_set_sbd0 i bd0_Card_order bd0s =  | 
| 75624 | 1432  | 
                map (fn thm => @{thm ordLess_ordIso_trans} OF
 | 
1433  | 
[bd0_Card_order RS mk_ordLess_csum n i thm, sbd0_ordIso]) bd0s;  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1434  | 
              val set_sbd0ss = @{map 3} mk_set_sbd0 ks bd0_Card_orders set_bd0ss;
 | 
| 56350 | 1435  | 
in  | 
| 75624 | 1436  | 
(lthy, sbd0, sbd0_card_order, sbd0_Cinfinite, sbd0_regularCard, set_sbd0ss)  | 
| 56350 | 1437  | 
end;  | 
1438  | 
||
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1439  | 
val (Ibnf_consts, lthy) =  | 
| 62324 | 1440  | 
          @{fold_map 9} (fn b => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn mapx =>
 | 
1441  | 
fn sets => fn wits => fn T => fn lthy =>  | 
|
| 56192 | 1442  | 
define_bnf_consts Hardly_Inline (user_policy Note_Some lthy) false (SOME deads)  | 
| 62324 | 1443  | 
map_b rel_b pred_b set_bs  | 
1444  | 
(((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd0), wits), NONE), NONE) lthy)  | 
|
1445  | 
bs map_bs rel_bs pred_bs set_bss fs_maps setss_by_bnf ctor_witss Ts lthy;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1446  | 
|
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1447  | 
val ((((((((((((((Izs, (Izs1, Izs1')), (Izs2, Izs2')), xFs), yFs))), Iphis), Ipsi1s),  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1448  | 
Ipsi2s), fs), fs_copy), us), (ys, ys')), _) =  | 
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1449  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1450  | 
|> mk_Frees "z" Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1451  | 
||>> mk_Frees' "z1" Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1452  | 
||>> mk_Frees' "z2" Ts'  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1453  | 
||>> mk_Frees "x" FTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1454  | 
||>> mk_Frees "y" FTs'  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1455  | 
||>> mk_Frees "R" IphiTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1456  | 
||>> mk_Frees "R" Ipsi1Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1457  | 
||>> mk_Frees "Q" Ipsi2Ts  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1458  | 
||>> mk_Frees "f" fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1459  | 
||>> mk_Frees "f" fTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1460  | 
||>> mk_Frees "u" uTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1461  | 
||>> mk_Frees' "y" passiveAs;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1462  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1463  | 
        val (_, Iconsts, Iconst_defs, mk_Iconsts) = @{split_list 4} Ibnf_consts;
 | 
| 62324 | 1464  | 
        val (_, Isetss, Ibds_Ds, Iwitss_Ds, _, _) = @{split_list 6} Iconsts;
 | 
1465  | 
val (Imap_defs, Iset_defss, Ibd_defs, Iwit_defss, Irel_defs, Ipred_defs) =  | 
|
1466  | 
          @{split_list 6} Iconst_defs;
 | 
|
1467  | 
val (mk_Imaps_Ds, mk_It_Ds, _, mk_Irels_Ds, mk_Ipreds_Ds, _, _) =  | 
|
1468  | 
          @{split_list 7} mk_Iconsts;
 | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1469  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
1470  | 
val Irel_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Irel_defs;  | 
| 
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67405 
diff
changeset
 | 
1471  | 
val Ipred_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Ipred_defs;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1472  | 
val Iset_defs = flat Iset_defss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1473  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1474  | 
fun mk_Imaps As Bs = map (fn mk => mk deads As Bs) mk_Imaps_Ds;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1475  | 
fun mk_Isetss As = map2 (fn mk => fn Isets => map (mk deads As) Isets) mk_It_Ds Isetss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1476  | 
val Ibds = map2 (fn mk => mk deads passiveAs) mk_It_Ds Ibds_Ds;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1477  | 
val Iwitss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1478  | 
map2 (fn mk => fn Iwits => map (mk deads passiveAs o snd) Iwits) mk_It_Ds Iwitss_Ds;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1479  | 
fun mk_Irels As Bs = map (fn mk => mk deads As Bs) mk_Irels_Ds;  | 
| 62324 | 1480  | 
fun mk_Ipreds As = map (fn mk => mk deads As) mk_Ipreds_Ds;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1481  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1482  | 
val Imaps = mk_Imaps passiveAs passiveBs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1483  | 
val fs_Imaps = map (fn m => Term.list_comb (m, fs)) Imaps;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1484  | 
val fs_copy_Imaps = map (fn m => Term.list_comb (m, fs_copy)) Imaps;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1485  | 
val (Isetss_by_range, Isetss_by_bnf) = `transpose (mk_Isetss passiveAs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1486  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1487  | 
val map_setss = map (fn T => map2 (fn Ds =>  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1488  | 
mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) passiveAs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1489  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1490  | 
val timer = time (timer "bnf constants for the new datatypes");  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1491  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1492  | 
val (ctor_Imap_thms, ctor_Imap_o_thms) =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1493  | 
let  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1494  | 
fun mk_goal fs_map map ctor ctor' =  | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1495  | 
mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor),  | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1496  | 
HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_Imaps)));  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1497  | 
            val goals = @{map 4} mk_goal fs_Imaps map_FTFT's ctors ctor's;
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1498  | 
val maps =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1499  | 
              @{map 4} (fn goal => fn foldx => fn map_comp_id => fn map_cong0 =>
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1500  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1501  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1502  | 
                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1503  | 
mk_map_tac ctxt m n foldx map_comp_id map_cong0))  | 
| 70494 | 1504  | 
|> Thm.close_derivation \<^here>)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1505  | 
goals ctor_fold_thms map_comp_id_thms map_cong0s;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1506  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1507  | 
            `(map (fn thm => thm RS @{thm comp_eq_dest})) maps
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1508  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1509  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1510  | 
val (ctor_Imap_unique_thms, ctor_Imap_unique_thm) =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1511  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1512  | 
fun mk_prem u map ctor ctor' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1513  | 
mk_Trueprop_eq (HOLogic.mk_comp (u, ctor),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1514  | 
HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ us)));  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1515  | 
            val prems = @{map 4} mk_prem us map_FTFT's ctors ctor's;
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1516  | 
val goal =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1517  | 
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1518  | 
(map2 (curry HOLogic.mk_eq) us fs_Imaps));  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1519  | 
val vars = fold (Variable.add_free_names lthy) (goal :: prems) [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1520  | 
val unique = Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal))  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1521  | 
              (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
 | 
| 55197 | 1522  | 
mk_ctor_map_unique_tac ctxt ctor_fold_unique_thm sym_map_comps)  | 
| 70494 | 1523  | 
|> Thm.close_derivation \<^here>;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1524  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1525  | 
`split_conj_thm unique  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1526  | 
end;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1527  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1528  | 
val timer = time (timer "map functions for the new datatypes");  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1529  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1530  | 
val ctor_Iset_thmss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1531  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1532  | 
fun mk_goal sets ctor set col map =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1533  | 
mk_Trueprop_eq (HOLogic.mk_comp (set, ctor),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1534  | 
HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets)));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1535  | 
val goalss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1536  | 
              @{map 3} (fn sets => @{map 4} (mk_goal sets) ctors sets)
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1537  | 
Isetss_by_range colss map_setss;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1538  | 
val setss = map (map2 (fn foldx => fn goal =>  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1539  | 
                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
 | 
| 60728 | 1540  | 
unfold_thms_tac ctxt Iset_defs THEN mk_set_tac ctxt foldx)  | 
| 70494 | 1541  | 
|> Thm.close_derivation \<^here>)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1542  | 
ctor_fold_thms) goalss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1543  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1544  | 
fun mk_simp_goal pas_set act_sets sets ctor z set =  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1545  | 
mk_Trueprop_eq (set $ (ctor $ z),  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1546  | 
mk_union (pas_set $ z,  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1547  | 
Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ z)) act_sets sets)));  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1548  | 
val simp_goalss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1549  | 
map2 (fn i => fn sets =>  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1550  | 
                @{map 4} (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets)
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1551  | 
FTs_setss ctors xFs sets)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1552  | 
ls Isetss_by_range;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1553  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1554  | 
            val ctor_setss = @{map 3} (fn i => @{map 3} (fn set_nats => fn goal => fn set =>
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1555  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1556  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 60728 | 1557  | 
                  (fn {context = ctxt, prems = _} =>
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1558  | 
mk_ctor_set_tac ctxt set (nth set_nats (i - 1)) (drop m set_nats)))  | 
| 70494 | 1559  | 
|> Thm.close_derivation \<^here>)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1560  | 
set_mapss) ls simp_goalss setss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1561  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1562  | 
ctor_setss  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1563  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1564  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1565  | 
        fun mk_set_thms ctor_set = (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper1}]) ::
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1566  | 
          map (fn i => (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper2}]) RS
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1567  | 
(mk_Un_upper n i RS subset_trans) RSN  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1568  | 
            (2, @{thm UN_upper} RS subset_trans))
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1569  | 
(1 upto n);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1570  | 
val set_Iset_thmsss = transpose (map (map mk_set_thms) ctor_Iset_thmss);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1571  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1572  | 
val timer = time (timer "set functions for the new datatypes");  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1573  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1574  | 
val cxs = map (SOME o Thm.cterm_of lthy) Izs;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1575  | 
val Isetss_by_range' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1576  | 
map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1577  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1578  | 
val Iset_Imap0_thmss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1579  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1580  | 
fun mk_set_map0 f map z set set' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1581  | 
HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1582  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1583  | 
fun mk_cphi f map z set set' = Thm.cterm_of lthy  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1584  | 
(Term.absfree (dest_Free z) (mk_set_map0 f map z set set'));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1585  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1586  | 
val csetss = map (map (Thm.cterm_of lthy)) Isetss_by_range';  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1587  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1588  | 
            val cphiss = @{map 3} (fn f => fn sets => fn sets' =>
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1589  | 
              (@{map 4} (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range';
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1590  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1591  | 
val inducts = map (fn cphis =>  | 
| 60801 | 1592  | 
Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1593  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1594  | 
val goals =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1595  | 
              @{map 3} (fn f => fn sets => fn sets' =>
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1596  | 
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1597  | 
                  (@{map 4} (mk_set_map0 f) fs_Imaps Izs sets sets')))
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1598  | 
fs Isetss_by_range Isetss_by_range';  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1599  | 
|
| 60728 | 1600  | 
fun mk_tac ctxt induct = mk_set_nat_tac ctxt m (rtac ctxt induct) set_mapss ctor_Imap_thms;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1601  | 
val thms =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1602  | 
              @{map 5} (fn goal => fn csets => fn ctor_sets => fn induct => fn i =>
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1603  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1604  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1605  | 
                  (fn {context = ctxt, prems = _} => mk_tac ctxt induct csets ctor_sets i))
 | 
| 70494 | 1606  | 
|> Thm.close_derivation \<^here>)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1607  | 
goals csetss ctor_Iset_thmss inducts ls;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1608  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1609  | 
map split_conj_thm thms  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1610  | 
end;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1611  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1612  | 
val Iset_bd_thmss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1613  | 
let  | 
| 75624 | 1614  | 
fun mk_set_bd z bd set = mk_ordLess (mk_card_of (set $ z)) bd;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1615  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1616  | 
fun mk_cphi z set = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_bd z sbd0 set));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1617  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1618  | 
val cphiss = map (map2 mk_cphi Izs) Isetss_by_range;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1619  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1620  | 
val inducts = map (fn cphis =>  | 
| 60801 | 1621  | 
Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1622  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1623  | 
val goals =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1624  | 
map (fn sets =>  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1625  | 
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1626  | 
                  (@{map 3} mk_set_bd Izs Ibds sets))) Isetss_by_range;
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1627  | 
|
| 75624 | 1628  | 
fun mk_tac ctxt induct = mk_set_bd_tac ctxt m (rtac ctxt induct) sbd0_Cinfinite  | 
1629  | 
sbd0_regularCard set_sbd0ss;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1630  | 
val thms =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1631  | 
              @{map 4} (fn goal => fn ctor_sets => fn induct => fn i =>
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1632  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1633  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1634  | 
                    (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Ibd_defs THEN
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1635  | 
mk_tac ctxt induct ctor_sets i))  | 
| 70494 | 1636  | 
|> Thm.close_derivation \<^here>)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1637  | 
goals ctor_Iset_thmss inducts ls;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1638  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1639  | 
map split_conj_thm thms  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1640  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1641  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1642  | 
val Imap_cong0_thms =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1643  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1644  | 
fun mk_prem z set f g y y' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1645  | 
mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1646  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1647  | 
fun mk_map_cong0 sets z fmap gmap =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1648  | 
HOLogic.mk_imp  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1649  | 
                (Library.foldr1 HOLogic.mk_conj (@{map 5} (mk_prem z) sets fs fs_copy ys ys'),
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1650  | 
HOLogic.mk_eq (fmap $ z, gmap $ z));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1651  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1652  | 
fun mk_cphi sets z fmap gmap =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1653  | 
Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap));  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1654  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1655  | 
            val cphis = @{map 4} mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps;
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1656  | 
|
| 60801 | 1657  | 
val induct = Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1658  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1659  | 
val goal =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1660  | 
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1661  | 
                (@{map 4} mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps));
 | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1662  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1663  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1664  | 
val thm = Goal.prove_sorry lthy vars [] goal  | 
| 60728 | 1665  | 
                (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt (rtac ctxt induct) set_Iset_thmsss
 | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1666  | 
map_cong0s ctor_Imap_thms)  | 
| 70494 | 1667  | 
|> Thm.close_derivation \<^here>;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1668  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1669  | 
split_conj_thm thm  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1670  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1671  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51869 
diff
changeset
 | 
1672  | 
val in_rels = map in_rel_of_bnf bnfs;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1673  | 
        val in_Irels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1674  | 
Irel_unabs_defs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1675  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1676  | 
val ctor_Iset_incl_thmss = map (map hd) set_Iset_thmsss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1677  | 
val ctor_set_Iset_incl_thmsss = map (transpose o map tl) set_Iset_thmsss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1678  | 
val ctor_Iset_thmss' = transpose ctor_Iset_thmss;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1679  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1680  | 
val Irels = mk_Irels passiveAs passiveBs;  | 
| 62324 | 1681  | 
val Ipreds = mk_Ipreds passiveAs;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1682  | 
val Irelphis = map (fn rel => Term.list_comb (rel, Iphis)) Irels;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1683  | 
val relphis = map (fn rel => Term.list_comb (rel, Iphis @ Irelphis)) rels;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1684  | 
val Irelpsi1s = map (fn rel => Term.list_comb (rel, Ipsi1s)) (mk_Irels passiveAs passiveCs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1685  | 
val Irelpsi2s = map (fn rel => Term.list_comb (rel, Ipsi2s)) (mk_Irels passiveCs passiveBs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1686  | 
val Irelpsi12s = map (fn rel =>  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1687  | 
Term.list_comb (rel, map2 (curry mk_rel_compp) Ipsi1s Ipsi2s)) Irels;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1688  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51869 
diff
changeset
 | 
1689  | 
val ctor_Irel_thms =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1690  | 
let  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1691  | 
fun mk_goal xF yF ctor ctor' Irelphi relphi =  | 
| 
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1692  | 
mk_Trueprop_eq (Irelphi $ (ctor $ xF) $ (ctor' $ yF), relphi $ xF $ yF);  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1693  | 
            val goals = @{map 6} mk_goal xFs yFs ctors ctor's Irelphis relphis;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1694  | 
in  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1695  | 
            @{map 12} (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 =>
 | 
| 
49542
 
b39354db8629
renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49541 
diff
changeset
 | 
1696  | 
fn ctor_map => fn ctor_sets => fn ctor_inject => fn ctor_dtor =>  | 
| 53289 | 1697  | 
fn set_map0s => fn ctor_set_incls => fn ctor_set_set_inclss =>  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1698  | 
Variable.add_free_names lthy goal []  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1699  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
| 61271 | 1700  | 
               (fn {context = ctxt, prems = _} =>
 | 
1701  | 
mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1702  | 
ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss))  | 
| 70494 | 1703  | 
|> Thm.close_derivation \<^here>)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1704  | 
ks goals in_rels map_comps map_cong0s ctor_Imap_thms ctor_Iset_thmss'  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1705  | 
ctor_inject_thms ctor_dtor_thms set_mapss ctor_Iset_incl_thmss  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1706  | 
ctor_set_Iset_incl_thmsss  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1707  | 
end;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1708  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1709  | 
val le_Irel_OO_thm =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1710  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1711  | 
fun mk_le_Irel_OO Irelpsi1 Irelpsi2 Irelpsi12 Iz1 Iz2 =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1712  | 
HOLogic.mk_imp (mk_rel_compp (Irelpsi1, Irelpsi2) $ Iz1 $ Iz2,  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1713  | 
Irelpsi12 $ Iz1 $ Iz2);  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1714  | 
            val goals = @{map 5} mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2;
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1715  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1716  | 
val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct2_params;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1717  | 
val cxs = map (SOME o Thm.cterm_of lthy) (splice Izs1 Izs2);  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59580 
diff
changeset
 | 
1718  | 
fun mk_cphi z1 z2 goal = SOME (Thm.cterm_of lthy (Term.absfree z1 (Term.absfree z2 goal)));  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
1719  | 
            val cphis = @{map 3} mk_cphi Izs1' Izs2' goals;
 | 
| 60801 | 1720  | 
val induct = Thm.instantiate' cTs (cphis @ cxs) ctor_induct2_thm;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1721  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1722  | 
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1723  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1724  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61272 
diff
changeset
 | 
1725  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
56272
 
159c07ceb18c
prove theorems with fixed variables (export afterwards)
 
traytel 
parents: 
56263 
diff
changeset
 | 
1726  | 
              (fn {context = ctxt, prems = _} => mk_le_rel_OO_tac ctxt m induct ctor_nchotomy_thms
 | 
| 57967 | 1727  | 
ctor_Irel_thms rel_mono_strong0s le_rel_OOs)  | 
| 70494 | 1728  | 
|> Thm.close_derivation \<^here>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1729  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1730  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1731  | 
val timer = time (timer "helpers for BNF properties");  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1732  | 
|
| 60728 | 1733  | 
val map_id0_tacs = map (fn thm => fn ctxt => mk_map_id0_tac ctxt map_id0s thm)  | 
1734  | 
ctor_Imap_unique_thms;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1735  | 
val map_comp0_tacs =  | 
| 60728 | 1736  | 
map2 (fn thm => fn i => fn ctxt =>  | 
1737  | 
mk_map_comp0_tac ctxt map_comps ctor_Imap_thms thm i)  | 
|
1738  | 
ctor_Imap_unique_thms ks;  | 
|
| 55197 | 1739  | 
val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) Imap_cong0_thms;  | 
| 60728 | 1740  | 
val set_map0_tacss = map (map (fn thm => fn ctxt => mk_set_map0_tac ctxt thm))  | 
1741  | 
(transpose Iset_Imap0_thmss);  | 
|
| 55197 | 1742  | 
val bd_co_tacs = replicate n (fn ctxt =>  | 
| 60728 | 1743  | 
unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt sbd0_card_order 1);  | 
| 55197 | 1744  | 
val bd_cinf_tacs = replicate n (fn ctxt =>  | 
| 60728 | 1745  | 
unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt (sbd0_Cinfinite RS conjunct1) 1);  | 
| 75624 | 1746  | 
val bd_reg_tacs = replicate n (fn ctxt =>  | 
1747  | 
unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt sbd0_regularCard 1);  | 
|
| 60728 | 1748  | 
val set_bd_tacss = map (map (fn thm => fn ctxt => rtac ctxt thm 1)) (transpose Iset_bd_thmss);  | 
1749  | 
val le_rel_OO_tacs = map (fn i => fn ctxt =>  | 
|
1750  | 
          (rtac ctxt @{thm predicate2I} THEN' etac ctxt (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1) ks;
 | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1751  | 
|
| 60728 | 1752  | 
val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Irel_unabs_defs;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1753  | 
|
| 62324 | 1754  | 
val pred_set_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Ipred_unabs_defs;  | 
1755  | 
||
| 75624 | 1756  | 
        val tacss = @{map 11} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss
 | 
1757  | 
bd_co_tacs bd_cinf_tacs bd_reg_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs pred_set_tacs;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1758  | 
|
| 55197 | 1759  | 
fun wit_tac ctxt = unfold_thms_tac ctxt (flat Iwit_defss) THEN  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1760  | 
mk_wit_tac ctxt n (flat ctor_Iset_thmss) (maps wit_thms_of_bnf bnfs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1761  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1762  | 
val (Ibnfs, lthy) =  | 
| 62324 | 1763  | 
          @{fold_map 6} (fn tacs => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn consts =>
 | 
| 
56016
 
8875cdcfc85b
unfold intermediate definitions after sealing the bnf
 
traytel 
parents: 
55901 
diff
changeset
 | 
1764  | 
bnf_def Do_Inline (user_policy Note_Some) false I tacs wit_tac (SOME deads)  | 
| 62324 | 1765  | 
map_b rel_b pred_b set_bs consts)  | 
1766  | 
tacss map_bs rel_bs pred_bs set_bss  | 
|
1767  | 
(((((((replicate n Binding.empty ~~ Ts) ~~ Imaps) ~~ Isetss_by_bnf) ~~ Ibds) ~~  | 
|
1768  | 
Iwitss) ~~ map SOME Irels) ~~ map SOME Ipreds) lthy;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1769  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1770  | 
val timer = time (timer "registered new datatypes as BNFs");  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1771  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1772  | 
val ls' = if m = 1 then [0] else ls  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1773  | 
|
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1774  | 
val Ibnf_common_notes =  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1775  | 
[(ctor_map_uniqueN, [ctor_Imap_unique_thm])]  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1776  | 
|> map (fn (thmN, thms) =>  | 
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1777  | 
((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));  | 
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1778  | 
|
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1779  | 
val Ibnf_notes =  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1780  | 
[(ctor_mapN, map single ctor_Imap_thms),  | 
| 49541 | 1781  | 
(ctor_relN, map single ctor_Irel_thms),  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1782  | 
(ctor_set_inclN, ctor_Iset_incl_thmss),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1783  | 
(ctor_set_set_inclN, map flat ctor_set_Iset_incl_thmsss)] @  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54793 
diff
changeset
 | 
1784  | 
map2 (fn i => fn thms => (mk_ctor_setN i, map single thms)) ls' ctor_Iset_thmss  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1785  | 
|> maps (fn (thmN, thmss) =>  | 
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1786  | 
map2 (fn b => fn thms =>  | 
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1787  | 
((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))  | 
| 
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1788  | 
bs thmss)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1789  | 
in  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1790  | 
(timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Imap_unique_thm, ctor_Iset_thmss',  | 
| 53567 | 1791  | 
ctor_Irel_thms, Ibnf_common_notes @ Ibnf_notes, lthy)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1792  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1793  | 
|
| 
61272
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1794  | 
val ((((((xFs, yFs)), Iphis), activephis), activeIphis), _) =  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1795  | 
lthy  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1796  | 
|> mk_Frees "x" FTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1797  | 
||>> mk_Frees "y" FTs'  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1798  | 
||>> mk_Frees "R" IphiTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1799  | 
||>> mk_Frees "S" activephiTs  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1800  | 
||>> mk_Frees "IR" activeIphiTs;  | 
| 
 
f49644098959
restructure fresh variable generation to make exports more wellformed
 
traytel 
parents: 
61271 
diff
changeset
 | 
1801  | 
|
| 62827 | 1802  | 
val ctor_fold_o_Imap_thms = mk_xtor_co_iter_o_map_thms Least_FP false m ctor_fold_unique_thm  | 
| 64413 | 1803  | 
ctor_Imap_o_thms (map (mk_pointfree2 lthy) ctor_fold_thms) sym_map_comps map_cong0s;  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52911 
diff
changeset
 | 
1804  | 
|
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1805  | 
val Irels = if m = 0 then map HOLogic.eq_const Ts  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1806  | 
else map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs;  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1807  | 
val Irel_induct_thm =  | 
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58578 
diff
changeset
 | 
1808  | 
mk_xtor_rel_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1809  | 
        (fn {context = ctxt, prems = IHs} => mk_rel_induct_tac ctxt IHs m ctor_induct2_thm ks
 | 
| 57967 | 1810  | 
ctor_Irel_thms rel_mono_strong0s) lthy;  | 
| 51918 | 1811  | 
|
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1812  | 
val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1813  | 
val ctor_fold_transfer_thms =  | 
| 62827 | 1814  | 
mk_xtor_co_iter_transfer_thms Least_FP rels activephis activephis Irels Iphis  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1815  | 
(mk_folds passiveAs activeAs) (mk_folds passiveBs activeBs)  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1816  | 
        (fn {context = ctxt, prems = _} => mk_fold_transfer_tac ctxt m Irel_induct_thm
 | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1817  | 
(map map_transfer_of_bnf bnfs) ctor_fold_thms)  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1818  | 
lthy;  | 
| 52731 | 1819  | 
|
| 62905 | 1820  | 
val timer = time (timer "relator induction");  | 
| 58443 | 1821  | 
|
| 62905 | 1822  | 
fun mk_Ts As = map (typ_subst_atomic (passiveAs ~~ As)) Ts;  | 
1823  | 
val export = map (Morphism.term (Local_Theory.target_morphism lthy))  | 
|
1824  | 
val ((recs, (ctor_rec_thms, ctor_rec_unique_thm, ctor_rec_o_Imap_thms, ctor_rec_transfer_thms)),  | 
|
1825  | 
lthy) = lthy  | 
|
1826  | 
|> derive_xtor_co_recs Least_FP external_bs mk_Ts (Dss, resDs) bnfs  | 
|
1827  | 
(export ctors) (export folds)  | 
|
| 63045 | 1828  | 
ctor_fold_unique_thm ctor_fold_thms ctor_fold_transfer_thms ctor_Imap_thms ctor_Irel_thms  | 
1829  | 
(replicate n NONE);  | 
|
| 62905 | 1830  | 
|
1831  | 
val timer = time (timer "recursor");  | 
|
| 51918 | 1832  | 
|
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1833  | 
val common_notes =  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1834  | 
[(ctor_inductN, [ctor_induct_thm]),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1835  | 
(ctor_induct2N, [ctor_induct2_thm]),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1836  | 
(ctor_rel_inductN, [Irel_induct_thm])]  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1837  | 
|> map (fn (thmN, thms) =>  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1838  | 
((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49074 
diff
changeset
 | 
1839  | 
|
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1840  | 
val notes =  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1841  | 
[(ctor_dtorN, ctor_dtor_thms),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1842  | 
(ctor_exhaustN, ctor_exhaust_thms),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1843  | 
(ctor_foldN, ctor_fold_thms),  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
1844  | 
(ctor_fold_o_mapN, ctor_fold_o_Imap_thms),  | 
| 
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
1845  | 
(ctor_fold_transferN, ctor_fold_transfer_thms),  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1846  | 
(ctor_fold_uniqueN, ctor_fold_unique_thms),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1847  | 
(ctor_injectN, ctor_inject_thms),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1848  | 
(dtor_ctorN, dtor_ctor_thms),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1849  | 
(dtor_exhaustN, dtor_exhaust_thms),  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1850  | 
(dtor_injectN, dtor_inject_thms)]  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1851  | 
|> map (apsnd (map single))  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1852  | 
|> maps (fn (thmN, thmss) =>  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1853  | 
map2 (fn b => fn thms =>  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1854  | 
((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))  | 
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1855  | 
bs thmss);  | 
| 
53568
 
f9456284048f
conceal low-level noted facts (+ FIXME to get rid of the notes altogether eventually)
 
traytel 
parents: 
53567 
diff
changeset
 | 
1856  | 
|
| 62093 | 1857  | 
val lthy' = lthy |> internals ? snd o Local_Theory.notes (common_notes @ notes @ Ibnf_notes);  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1858  | 
|
| 
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1859  | 
val fp_res =  | 
| 
62684
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62324 
diff
changeset
 | 
1860  | 
      {Ts = Ts, bnfs = Ibnfs, pre_bnfs = bnfs, absT_infos = absT_infos,
 | 
| 62907 | 1861  | 
ctors = ctors, dtors = dtors, xtor_un_folds = folds, xtor_co_recs = export recs,  | 
| 
62684
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62324 
diff
changeset
 | 
1862  | 
xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms,  | 
| 59819 | 1863  | 
ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
1864  | 
dtor_injects = dtor_inject_thms, xtor_maps = ctor_Imap_thms,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1865  | 
xtor_map_unique = ctor_Imap_unique_thm, xtor_setss = ctor_Iset_thmss',  | 
| 62907 | 1866  | 
xtor_rels = ctor_Irel_thms, xtor_un_fold_thms = ctor_fold_thms,  | 
1867  | 
xtor_co_rec_thms = ctor_rec_thms, xtor_un_fold_unique = ctor_fold_unique_thm,  | 
|
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
1868  | 
xtor_co_rec_unique = ctor_rec_unique_thm,  | 
| 62907 | 1869  | 
xtor_un_fold_o_maps = ctor_fold_o_Imap_thms,  | 
| 62721 | 1870  | 
xtor_co_rec_o_maps = ctor_rec_o_Imap_thms,  | 
| 62907 | 1871  | 
xtor_un_fold_transfers = ctor_fold_transfer_thms,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
1872  | 
xtor_co_rec_transfers = ctor_rec_transfer_thms, xtor_rel_co_induct = Irel_induct_thm,  | 
| 
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
1873  | 
dtor_set_inducts = []};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1874  | 
in  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57567 
diff
changeset
 | 
1875  | 
timer; (fp_res, lthy')  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1876  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1877  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1878  | 
val _ =  | 
| 69593 | 1879  | 
Outer_Syntax.local_theory \<^command_keyword>\<open>datatype\<close> "define inductive datatypes"  | 
| 
58305
 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 
blanchet 
parents: 
58256 
diff
changeset
 | 
1880  | 
(parse_co_datatype_cmd Least_FP construct_lfp);  | 
| 
 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 
blanchet 
parents: 
58256 
diff
changeset
 | 
1881  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1882  | 
end;  |