| author | traytel | 
| Thu, 06 Mar 2014 14:14:54 +0100 | |
| changeset 55935 | 8f20d09d294e | 
| parent 55930 | 25a90cebbbe5 | 
| child 55937 | 18e52e8c6300 | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_comp.ML  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
Copyright 2012  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
Composition of bounded natural functors.  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
|
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
9  | 
val inline_ref = Unsynchronized.ref true;  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
10  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
11  | 
signature BNF_COMP =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
12  | 
sig  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51782 
diff
changeset
 | 
13  | 
val ID_bnf: BNF_Def.bnf  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51782 
diff
changeset
 | 
14  | 
val DEADID_bnf: BNF_Def.bnf  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49538 
diff
changeset
 | 
15  | 
|
| 55706 | 16  | 
type comp_cache  | 
| 49502 | 17  | 
type unfold_set  | 
| 55706 | 18  | 
|
19  | 
val empty_comp_cache: comp_cache  | 
|
| 49502 | 20  | 
val empty_unfolds: unfold_set  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
21  | 
|
| 53222 | 22  | 
exception BAD_DEAD of typ * typ  | 
23  | 
||
| 
55856
 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 
blanchet 
parents: 
55855 
diff
changeset
 | 
24  | 
val bnf_of_typ: BNF_Def.inline_policy -> (binding -> binding) ->  | 
| 
55703
 
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
 
blanchet 
parents: 
55480 
diff
changeset
 | 
25  | 
((string * sort) list list -> (string * sort) list) -> (string * sort) list ->  | 
| 55904 | 26  | 
(string * sort) list -> typ -> (comp_cache * unfold_set) * local_theory ->  | 
27  | 
(BNF_Def.bnf * (typ list * typ list)) * ((comp_cache * unfold_set) * local_theory)  | 
|
| 49014 | 28  | 
val default_comp_sort: (string * sort) list list -> (string * sort) list  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
29  | 
val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->  | 
| 55904 | 30  | 
    (''a list list -> ''a list) -> BNF_Def.bnf list -> (comp_cache * unfold_set) * local_theory ->
 | 
31  | 
(int list list * ''a list) * (BNF_Def.bnf list * ((comp_cache * unfold_set) * local_theory))  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
32  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
33  | 
type absT_info =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
34  | 
    {absT: typ,
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
35  | 
repT: typ,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
36  | 
abs: term,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
37  | 
rep: term,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
38  | 
abs_inject: thm,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
39  | 
abs_inverse: thm,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
40  | 
type_definition: thm}  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
41  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
42  | 
val morph_absT_info: morphism -> absT_info -> absT_info  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
43  | 
val mk_absT: theory -> typ -> typ -> typ -> typ  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
44  | 
val mk_repT: typ -> typ -> typ -> typ  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
45  | 
val mk_abs: typ -> term -> term  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
46  | 
val mk_rep: typ -> term -> term  | 
| 53264 | 47  | 
val seal_bnf: (binding -> binding) -> unfold_set -> binding -> typ list -> BNF_Def.bnf ->  | 
| 55904 | 48  | 
local_theory -> (BNF_Def.bnf * (typ list * absT_info)) * local_theory  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
49  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
51  | 
structure BNF_Comp : BNF_COMP =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
52  | 
struct  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
54  | 
open BNF_Def  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
55  | 
open BNF_Util  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
56  | 
open BNF_Tactics  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
open BNF_Comp_Tactics  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
|
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
59  | 
val ID_bnf = the (bnf_of @{context} "BNF_Comp.ID");
 | 
| 
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
60  | 
val DEADID_bnf = the (bnf_of @{context} "BNF_Comp.DEADID");
 | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49538 
diff
changeset
 | 
61  | 
|
| 55706 | 62  | 
type comp_cache = (bnf * (typ list * typ list)) Typtab.table;  | 
63  | 
||
| 55904 | 64  | 
fun key_of_types s Ts = Type (s, Ts);  | 
65  | 
fun key_of_typess s = key_of_types s o map (key_of_types "");  | 
|
66  | 
fun typ_of_int n = Type (string_of_int n, []);  | 
|
67  | 
fun typ_of_bnf bnf =  | 
|
68  | 
key_of_typess "" [[T_of_bnf bnf], lives_of_bnf bnf, sort Term_Ord.typ_ord (deads_of_bnf bnf)];  | 
|
69  | 
||
70  | 
fun key_of_kill n bnf = key_of_types "k" [typ_of_int n, typ_of_bnf bnf];  | 
|
71  | 
fun key_of_lift n bnf = key_of_types "l" [typ_of_int n, typ_of_bnf bnf];  | 
|
72  | 
fun key_of_permute src dest bnf =  | 
|
73  | 
key_of_types "p" (map typ_of_int src @ map typ_of_int dest @ [typ_of_bnf bnf]);  | 
|
74  | 
fun key_of_compose oDs Dss Ass outer inners =  | 
|
75  | 
key_of_types "c" (map (key_of_typess "") [[oDs], Dss, Ass, [map typ_of_bnf (outer :: inners)]]);  | 
|
76  | 
||
77  | 
fun cache_comp_simple key cache (bnf, (unfold_set, lthy)) =  | 
|
78  | 
(bnf, ((Typtab.update (key, (bnf, ([], []))) cache, unfold_set), lthy));  | 
|
79  | 
||
80  | 
fun cache_comp key (bnf_Ds_As, ((cache, unfold_set), lthy)) =  | 
|
81  | 
(bnf_Ds_As, ((Typtab.update (key, bnf_Ds_As) cache, unfold_set), lthy));  | 
|
82  | 
||
| 55706 | 83  | 
(* TODO: Replace by "BNF_Defs.defs list"? *)  | 
| 49502 | 84  | 
type unfold_set = {
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
85  | 
map_unfolds: thm list,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
86  | 
set_unfoldss: thm list list,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
87  | 
rel_unfolds: thm list  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
88  | 
};  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
89  | 
|
| 55706 | 90  | 
val empty_comp_cache = Typtab.empty;  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
91  | 
val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = []};
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
92  | 
|
| 49503 | 93  | 
fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new;  | 
94  | 
fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms;  | 
|
95  | 
||
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
96  | 
fun add_to_unfolds map sets rel  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
97  | 
  {map_unfolds, set_unfoldss, rel_unfolds} =
 | 
| 49503 | 98  | 
  {map_unfolds = add_to_thms map_unfolds map,
 | 
99  | 
set_unfoldss = adds_to_thms set_unfoldss sets,  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
100  | 
rel_unfolds = add_to_thms rel_unfolds rel};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
101  | 
|
| 49503 | 102  | 
fun add_bnf_to_unfolds bnf =  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
103  | 
add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
105  | 
val bdTN = "bdT";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
106  | 
|
| 49425 | 107  | 
fun mk_killN n = "_kill" ^ string_of_int n;  | 
108  | 
fun mk_liftN n = "_lift" ^ string_of_int n;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
109  | 
fun mk_permuteN src dest =  | 
| 49425 | 110  | 
"_permute_" ^ implode (map string_of_int src) ^ "_" ^ implode (map string_of_int dest);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
111  | 
|
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
112  | 
val id_bnf_comp_def = @{thm id_bnf_comp_def}
 | 
| 
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
113  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
114  | 
(*copied from Envir.expand_term_free*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
115  | 
fun expand_term_const defs =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
116  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
117  | 
val eqs = map ((fn ((x, U), u) => (x, (U, u))) o apfst dest_Const) defs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
118  | 
val get = fn Const (x, _) => AList.lookup (op =) eqs x | _ => NONE;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
119  | 
in Envir.expand_term get end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
120  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
121  | 
fun is_sum_prod_natLeq (Const (@{const_name csum}, _) $ t $ u) = forall is_sum_prod_natLeq [t, u]
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
122  | 
  | is_sum_prod_natLeq (Const (@{const_name cprod}, _) $ t $ u) = forall is_sum_prod_natLeq [t, u]
 | 
| 55853 | 123  | 
  | is_sum_prod_natLeq t = t aconv @{term natLeq};
 | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
124  | 
|
| 49502 | 125  | 
fun clean_compose_bnf const_policy qualify b outer inners (unfold_set, lthy) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
126  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
127  | 
val olive = live_of_bnf outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
128  | 
val onwits = nwits_of_bnf outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
129  | 
val odead = dead_of_bnf outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
130  | 
val inner = hd inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
val ilive = live_of_bnf inner;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
132  | 
val ideads = map dead_of_bnf inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
133  | 
val inwitss = map nwits_of_bnf inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
134  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
135  | 
(* TODO: check olive = length inners > 0,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
136  | 
forall inner from inners. ilive = live,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
137  | 
forall inner from inners. idead = dead *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
139  | 
val (oDs, lthy1) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
140  | 
(Variable.invent_types (replicate odead HOLogic.typeS) lthy);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
141  | 
val (Dss, lthy2) = apfst (map (map TFree))  | 
| 55904 | 142  | 
(fold_map Variable.invent_types (map (fn n => replicate n HOLogic.typeS) ideads) lthy1);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
143  | 
val (Ass, lthy3) = apfst (replicate ilive o map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
144  | 
(Variable.invent_types (replicate ilive HOLogic.typeS) lthy2);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
145  | 
val As = if ilive > 0 then hd Ass else [];  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
146  | 
val Ass_repl = replicate olive As;  | 
| 55906 | 147  | 
val (Bs, names_lthy) = apfst (map TFree)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
148  | 
(Variable.invent_types (replicate ilive HOLogic.typeS) lthy3);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
149  | 
val Bss_repl = replicate olive Bs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
150  | 
|
| 55906 | 151  | 
val ((((fs', Qs'), Asets), xs), _) = names_lthy  | 
| 52923 | 152  | 
|> apfst snd o mk_Frees' "f" (map2 (curry op -->) As Bs)  | 
| 49463 | 153  | 
||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs)  | 
| 49456 | 154  | 
||>> mk_Frees "A" (map HOLogic.mk_setT As)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
155  | 
||>> mk_Frees "x" As;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
156  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
157  | 
val CAs = map3 mk_T_of_bnf Dss Ass_repl inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
158  | 
val CCA = mk_T_of_bnf oDs CAs outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
val CBs = map3 mk_T_of_bnf Dss Bss_repl inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
160  | 
val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
161  | 
val inner_setss = map3 mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
162  | 
val inner_bds = map3 mk_bd_of_bnf Dss Ass_repl inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
163  | 
val outer_bd = mk_bd_of_bnf oDs CAs outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
165  | 
(*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*)  | 
| 49303 | 166  | 
val mapx = fold_rev Term.abs fs'  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
167  | 
(Term.list_comb (mk_map_of_bnf oDs CAs CBs outer,  | 
| 49463 | 168  | 
map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
169  | 
mk_map_of_bnf Ds As Bs) Dss inners));  | 
| 49507 | 170  | 
(*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*)  | 
171  | 
val rel = fold_rev Term.abs Qs'  | 
|
172  | 
(Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer,  | 
|
| 49463 | 173  | 
map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o  | 
| 49507 | 174  | 
mk_rel_of_bnf Ds As Bs) Dss inners));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
176  | 
    (*Union o collect {outer.set_1 ... outer.set_m} o outer.map inner_1.set_i ... inner_m.set_i*)
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
177  | 
    (*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*)
 | 
| 49303 | 178  | 
fun mk_set i =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
179  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
180  | 
val (setTs, T) = `(replicate olive o HOLogic.mk_setT) (nth As i);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
181  | 
val outer_set = mk_collect  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
182  | 
(mk_sets_of_bnf (replicate olive oDs) (replicate olive setTs) outer)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
183  | 
(mk_T_of_bnf oDs setTs outer --> HOLogic.mk_setT T);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
184  | 
val inner_sets = map (fn sets => nth sets i) inner_setss;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
185  | 
val outer_map = mk_map_of_bnf oDs CAs setTs outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
186  | 
val map_inner_sets = Term.list_comb (outer_map, inner_sets);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
187  | 
val collect_image = mk_collect  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
188  | 
(map2 (fn f => fn set => HOLogic.mk_comp (mk_image f, set)) inner_sets outer_sets)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
189  | 
(CCA --> HOLogic.mk_setT T);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
190  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
191  | 
(Library.foldl1 HOLogic.mk_comp [mk_Union T, outer_set, map_inner_sets],  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
192  | 
HOLogic.mk_comp (mk_Union T, collect_image))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
193  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
194  | 
|
| 49303 | 195  | 
val (sets, sets_alt) = map_split mk_set (0 upto ilive - 1);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
196  | 
|
| 55906 | 197  | 
fun mk_simplified_set set =  | 
198  | 
let  | 
|
199  | 
val setT = fastype_of set;  | 
|
| 
55908
 
e6d570cb0f5e
no 'sorry' so that the schematic variable gets instantiated
 
blanchet 
parents: 
55906 
diff
changeset
 | 
200  | 
        val var_set' = Const (@{const_name id_bnf_comp}, setT --> setT) $ Var ((Name.uu, 0), setT);
 | 
| 
 
e6d570cb0f5e
no 'sorry' so that the schematic variable gets instantiated
 
blanchet 
parents: 
55906 
diff
changeset
 | 
201  | 
val goal = mk_Trueprop_eq (var_set', set);  | 
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55908 
diff
changeset
 | 
202  | 
        fun tac {context = ctxt, prems = _} =
 | 
| 
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55908 
diff
changeset
 | 
203  | 
mk_simplified_set_tac ctxt (collect_set_map_of_bnf outer);  | 
| 55906 | 204  | 
val set'_eq_set =  | 
| 
55908
 
e6d570cb0f5e
no 'sorry' so that the schematic variable gets instantiated
 
blanchet 
parents: 
55906 
diff
changeset
 | 
205  | 
Goal.prove names_lthy [] [] goal tac  | 
| 55906 | 206  | 
|> Thm.close_derivation;  | 
207  | 
val set' = fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of set'_eq_set)));  | 
|
208  | 
in  | 
|
209  | 
(set', set'_eq_set)  | 
|
210  | 
end;  | 
|
211  | 
||
212  | 
val (sets', set'_eq_sets) =  | 
|
213  | 
map_split mk_simplified_set sets  | 
|
214  | 
||> Proof_Context.export names_lthy lthy;  | 
|
215  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
216  | 
(*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*)  | 
| 54421 | 217  | 
val bd = mk_cprod (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
218  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
219  | 
val (bd', bd_ordIso_natLeq_thm_opt) =  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
220  | 
if is_sum_prod_natLeq bd then  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
221  | 
let  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
222  | 
          val bd' = @{term natLeq};
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
223  | 
val bd_bd' = HOLogic.mk_prod (bd, bd');  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
224  | 
          val ordIso = Const (@{const_name ordIso}, HOLogic.mk_setT (fastype_of bd_bd'));
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
225  | 
val goal = HOLogic.mk_Trueprop (HOLogic.mk_mem (bd_bd', ordIso));  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
226  | 
in  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
227  | 
(bd', SOME (Goal.prove_sorry lthy [] [] goal (K bd_ordIso_natLeq_tac)  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
228  | 
|> Thm.close_derivation))  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
229  | 
end  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
230  | 
else  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
231  | 
(bd, NONE);  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
232  | 
|
| 53270 | 233  | 
fun map_id0_tac _ =  | 
234  | 
mk_comp_map_id0_tac (map_id0_of_bnf outer) (map_cong0_of_bnf outer)  | 
|
235  | 
(map map_id0_of_bnf inners);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
236  | 
|
| 53287 | 237  | 
fun map_comp0_tac _ =  | 
238  | 
mk_comp_map_comp0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer)  | 
|
239  | 
(map map_comp0_of_bnf inners);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
240  | 
|
| 55906 | 241  | 
fun mk_single_set_map0_tac i ctxt =  | 
242  | 
mk_comp_set_map0_tac ctxt (nth set'_eq_sets i) (map_comp0_of_bnf outer)  | 
|
243  | 
(map_cong0_of_bnf outer) (collect_set_map_of_bnf outer)  | 
|
| 53289 | 244  | 
(map ((fn thms => nth thms i) o set_map0_of_bnf) inners);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
245  | 
|
| 53289 | 246  | 
val set_map0_tacs = map mk_single_set_map0_tac (0 upto ilive - 1);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
247  | 
|
| 49303 | 248  | 
fun bd_card_order_tac _ =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
249  | 
mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
250  | 
|
| 49303 | 251  | 
fun bd_cinfinite_tac _ =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
252  | 
mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
253  | 
|
| 49303 | 254  | 
val set_alt_thms =  | 
| 52059 | 255  | 
if Config.get lthy quick_and_dirty then  | 
| 49456 | 256  | 
[]  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
257  | 
else  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49019 
diff
changeset
 | 
258  | 
map (fn goal =>  | 
| 51551 | 259  | 
Goal.prove_sorry lthy [] [] goal  | 
| 49714 | 260  | 
            (fn {context = ctxt, prems = _} =>
 | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51761 
diff
changeset
 | 
261  | 
mk_comp_set_alt_tac ctxt (collect_set_map_of_bnf outer))  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49019 
diff
changeset
 | 
262  | 
|> Thm.close_derivation)  | 
| 49303 | 263  | 
(map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) sets sets_alt);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
264  | 
|
| 55906 | 265  | 
fun map_cong0_tac ctxt =  | 
266  | 
mk_comp_map_cong0_tac ctxt set'_eq_sets set_alt_thms (map_cong0_of_bnf outer)  | 
|
267  | 
(map map_cong0_of_bnf inners);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
268  | 
|
| 49303 | 269  | 
val set_bd_tacs =  | 
| 52059 | 270  | 
if Config.get lthy quick_and_dirty then  | 
| 49669 | 271  | 
replicate ilive (K all_tac)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
272  | 
else  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
273  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
274  | 
val outer_set_bds = set_bd_of_bnf outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
275  | 
val inner_set_bdss = map set_bd_of_bnf inners;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
276  | 
val inner_bd_Card_orders = map bd_Card_order_of_bnf inners;  | 
| 49303 | 277  | 
fun single_set_bd_thm i j =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
278  | 
            @{thm comp_single_set_bd} OF [nth inner_bd_Card_orders j, nth (nth inner_set_bdss j) i,
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
279  | 
nth outer_set_bds j]  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
280  | 
val single_set_bd_thmss =  | 
| 49303 | 281  | 
map ((fn f => map f (0 upto olive - 1)) o single_set_bd_thm) (0 upto ilive - 1);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
282  | 
in  | 
| 55906 | 283  | 
map3 (fn set'_eq_set => fn set_alt => fn single_set_bds => fn ctxt =>  | 
284  | 
mk_comp_set_bd_tac ctxt set'_eq_set bd_ordIso_natLeq_thm_opt set_alt single_set_bds)  | 
|
285  | 
set'_eq_sets set_alt_thms single_set_bd_thmss  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
286  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
287  | 
|
| 49303 | 288  | 
val in_alt_thm =  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
289  | 
let  | 
| 49303 | 290  | 
val inx = mk_in Asets sets CCA;  | 
291  | 
val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;  | 
|
292  | 
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));  | 
|
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
293  | 
in  | 
| 51551 | 294  | 
Goal.prove_sorry lthy [] [] goal  | 
| 49714 | 295  | 
          (fn {context = ctxt, prems = _} => mk_comp_in_alt_tac ctxt set_alt_thms)
 | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
296  | 
|> Thm.close_derivation  | 
| 
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
297  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
298  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
299  | 
fun le_rel_OO_tac _ = mk_le_rel_OO_tac (le_rel_OO_of_bnf outer) (rel_mono_of_bnf outer)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
300  | 
(map le_rel_OO_of_bnf inners);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
301  | 
|
| 55906 | 302  | 
fun rel_OO_Grp_tac ctxt =  | 
| 49456 | 303  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
304  | 
val outer_rel_Grp = rel_Grp_of_bnf outer RS sym;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
305  | 
val outer_rel_cong = rel_cong_of_bnf outer;  | 
| 49463 | 306  | 
val thm =  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
307  | 
          (trans OF [in_alt_thm RS @{thm OO_Grp_cong},
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
308  | 
             trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
309  | 
               [trans OF [outer_rel_Grp RS @{thm arg_cong[of _ _ conversep]},
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
310  | 
rel_conversep_of_bnf outer RS sym], outer_rel_Grp],  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
311  | 
trans OF [rel_OO_of_bnf outer RS sym, outer_rel_cong OF  | 
| 55906 | 312  | 
(map (fn bnf => rel_OO_Grp_of_bnf bnf RS sym) inners)]]] RS sym);  | 
| 49456 | 313  | 
in  | 
| 55906 | 314  | 
unfold_thms_tac ctxt set'_eq_sets THEN rtac thm 1  | 
| 49463 | 315  | 
end;  | 
| 49456 | 316  | 
|
| 53289 | 317  | 
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
318  | 
bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
319  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
320  | 
val outer_wits = mk_wits_of_bnf (replicate onwits oDs) (replicate onwits CAs) outer;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
321  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
322  | 
val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
323  | 
(map3 (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
324  | 
Dss inwitss inners);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
325  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
326  | 
val inner_witsss = map (map (nth inner_witss) o fst) outer_wits;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
327  | 
|
| 49303 | 328  | 
val wits = (inner_witsss, (map (single o snd) outer_wits))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
329  | 
|-> map2 (fold (map_product (fn iwit => fn owit => owit $ iwit)))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
330  | 
|> flat  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
331  | 
|> map (`(fn t => Term.add_frees t []))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
332  | 
|> minimize_wits  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
333  | 
|> map (fn (frees, t) => fold absfree frees t);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
334  | 
|
| 55197 | 335  | 
fun wit_tac ctxt =  | 
| 55906 | 336  | 
mk_comp_wit_tac ctxt set'_eq_sets (wit_thms_of_bnf outer) (collect_set_map_of_bnf outer)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
337  | 
(maps wit_thms_of_bnf inners);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
338  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
339  | 
val (bnf', lthy') =  | 
| 51758 | 340  | 
bnf_def const_policy (K Dont_Note) qualify tacs wit_tac (SOME (oDs @ flat Dss)) Binding.empty  | 
| 55906 | 341  | 
Binding.empty [] ((((((b, CCA), mapx), sets'), bd'), wits), SOME rel) lthy;  | 
342  | 
||
343  | 
val phi =  | 
|
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
344  | 
Morphism.thm_morphism "BNF" (unfold_thms lthy' [id_bnf_comp_def])  | 
| 55906 | 345  | 
$> Morphism.term_morphism "BNF" (Raw_Simplifier.rewrite_term (Proof_Context.theory_of lthy')  | 
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
346  | 
[id_bnf_comp_def] []);  | 
| 55906 | 347  | 
|
348  | 
val bnf'' = morph_bnf phi bnf';  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
349  | 
in  | 
| 55906 | 350  | 
(bnf'', (add_bnf_to_unfolds bnf'' unfold_set, lthy'))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
351  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
352  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
353  | 
(* Killing live variables *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
354  | 
|
| 55904 | 355  | 
fun raw_kill_bnf qualify n bnf (accum as (unfold_set, lthy)) =  | 
356  | 
if n = 0 then (bnf, accum) else  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
357  | 
let  | 
| 49425 | 358  | 
val b = Binding.suffix_name (mk_killN n) (name_of_bnf bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
359  | 
val live = live_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
360  | 
val dead = dead_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
361  | 
val nwits = nwits_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
362  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
363  | 
(* TODO: check 0 < n <= live *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
364  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
365  | 
val (Ds, lthy1) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
366  | 
(Variable.invent_types (replicate dead HOLogic.typeS) lthy);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
367  | 
val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
368  | 
(Variable.invent_types (replicate live HOLogic.typeS) lthy1);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
369  | 
val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
370  | 
(Variable.invent_types (replicate (live - n) HOLogic.typeS) lthy2);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
371  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
372  | 
val ((Asets, lives), _(*names_lthy*)) = lthy  | 
| 49456 | 373  | 
|> mk_Frees "A" (map HOLogic.mk_setT (drop n As))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
374  | 
||>> mk_Frees "x" (drop n As);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
375  | 
    val xs = map (fn T => HOLogic.choice_const T $ absdummy T @{term True}) killedAs @ lives;
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
376  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
377  | 
val T = mk_T_of_bnf Ds As bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
378  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
379  | 
(*bnf.map id ... id*)  | 
| 49303 | 380  | 
val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);  | 
| 49507 | 381  | 
(*bnf.rel (op =) ... (op =)*)  | 
382  | 
val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
383  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
384  | 
val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;  | 
| 49303 | 385  | 
val sets = drop n bnf_sets;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
386  | 
|
| 
55707
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
387  | 
val bd = mk_bd_of_bnf Ds As bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
388  | 
|
| 53270 | 389  | 
fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;  | 
| 55197 | 390  | 
fun map_comp0_tac ctxt =  | 
| 55067 | 391  | 
unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) ::  | 
392  | 
        @{thms comp_assoc id_comp comp_id}) THEN rtac refl 1;
 | 
|
| 55197 | 393  | 
fun map_cong0_tac ctxt =  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
394  | 
mk_kill_map_cong0_tac ctxt n (live - n) (map_cong0_of_bnf bnf);  | 
| 53289 | 395  | 
val set_map0_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_map0_of_bnf bnf));  | 
| 
55707
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
396  | 
fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;  | 
| 
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
397  | 
fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;  | 
| 
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
398  | 
val set_bd_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_bd_of_bnf bnf));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
399  | 
|
| 49303 | 400  | 
val in_alt_thm =  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
401  | 
let  | 
| 49303 | 402  | 
val inx = mk_in Asets sets T;  | 
403  | 
val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;  | 
|
404  | 
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));  | 
|
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
405  | 
in  | 
| 51551 | 406  | 
Goal.prove_sorry lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
407  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
408  | 
|
| 55197 | 409  | 
fun le_rel_OO_tac ctxt =  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
410  | 
      EVERY' [rtac @{thm ord_le_eq_trans}, rtac (le_rel_OO_of_bnf bnf)] 1 THEN
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
411  | 
      unfold_thms_tac ctxt @{thms eq_OO} THEN rtac refl 1;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
412  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
413  | 
fun rel_OO_Grp_tac _ =  | 
| 49456 | 414  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
415  | 
val rel_Grp = rel_Grp_of_bnf bnf RS sym  | 
| 49463 | 416  | 
val thm =  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
417  | 
          (trans OF [in_alt_thm RS @{thm OO_Grp_cong},
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
418  | 
            trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
419  | 
              [trans OF [rel_Grp RS @{thm arg_cong[of _ _ conversep]},
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
420  | 
rel_conversep_of_bnf bnf RS sym], rel_Grp],  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
421  | 
trans OF [rel_OO_of_bnf bnf RS sym, rel_cong_of_bnf bnf OF  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
422  | 
                (replicate n @{thm trans[OF Grp_UNIV_id[OF refl] eq_alt[symmetric]]} @
 | 
| 52660 | 423  | 
                 replicate (live - n) @{thm Grp_fst_snd})]]] RS sym);
 | 
| 49456 | 424  | 
in  | 
| 49463 | 425  | 
rtac thm 1  | 
| 49456 | 426  | 
end;  | 
427  | 
||
| 53289 | 428  | 
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
429  | 
bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
430  | 
|
| 49303 | 431  | 
val bnf_wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
432  | 
|
| 49303 | 433  | 
val wits = map (fn t => fold absfree (Term.add_frees t []) t)  | 
434  | 
(map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
435  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
436  | 
fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
437  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
438  | 
val (bnf', lthy') =  | 
| 51758 | 439  | 
bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME (killedAs @ Ds)) Binding.empty  | 
| 54421 | 440  | 
Binding.empty [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
441  | 
in  | 
| 49503 | 442  | 
(bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
443  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
444  | 
|
| 55904 | 445  | 
fun kill_bnf qualify n bnf (accum as ((cache, unfold_set), lthy)) =  | 
446  | 
let val key = key_of_kill n bnf in  | 
|
447  | 
(case Typtab.lookup cache key of  | 
|
448  | 
SOME (bnf, _) => (bnf, accum)  | 
|
449  | 
| NONE => cache_comp_simple key cache (raw_kill_bnf qualify n bnf (unfold_set, lthy)))  | 
|
450  | 
end;  | 
|
451  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
452  | 
(* Adding dummy live variables *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
453  | 
|
| 55904 | 454  | 
fun raw_lift_bnf qualify n bnf (accum as (unfold_set, lthy)) =  | 
455  | 
if n = 0 then (bnf, accum) else  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
456  | 
let  | 
| 49425 | 457  | 
val b = Binding.suffix_name (mk_liftN n) (name_of_bnf bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
458  | 
val live = live_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
459  | 
val dead = dead_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
460  | 
val nwits = nwits_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
461  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
462  | 
(* TODO: check 0 < n *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
463  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
464  | 
val (Ds, lthy1) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
465  | 
(Variable.invent_types (replicate dead HOLogic.typeS) lthy);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
466  | 
val ((newAs, As), lthy2) = apfst (chop n o map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
467  | 
(Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy1);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
468  | 
val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
469  | 
(Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy2);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
471  | 
val (Asets, _(*names_lthy*)) = lthy  | 
| 49456 | 472  | 
|> mk_Frees "A" (map HOLogic.mk_setT (newAs @ As));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
473  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
474  | 
val T = mk_T_of_bnf Ds As bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
475  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
476  | 
(*%f1 ... fn. bnf.map*)  | 
| 49303 | 477  | 
val mapx =  | 
| 52923 | 478  | 
fold_rev Term.absdummy (map2 (curry op -->) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);  | 
| 49507 | 479  | 
(*%Q1 ... Qn. bnf.rel*)  | 
480  | 
val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
481  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
482  | 
val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;  | 
| 49303 | 483  | 
val sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
484  | 
|
| 49303 | 485  | 
val bd = mk_bd_of_bnf Ds As bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
486  | 
|
| 53270 | 487  | 
fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;  | 
| 55197 | 488  | 
fun map_comp0_tac ctxt =  | 
| 55067 | 489  | 
unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) ::  | 
490  | 
        @{thms comp_assoc id_comp comp_id}) THEN rtac refl 1;
 | 
|
| 55197 | 491  | 
fun map_cong0_tac ctxt =  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
492  | 
rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);  | 
| 53289 | 493  | 
val set_map0_tacs =  | 
| 52059 | 494  | 
if Config.get lthy quick_and_dirty then  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
495  | 
replicate (n + live) (K all_tac)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
496  | 
else  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
497  | 
replicate n (K empty_natural_tac) @  | 
| 53289 | 498  | 
map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf);  | 
| 49303 | 499  | 
fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;  | 
500  | 
fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;  | 
|
501  | 
val set_bd_tacs =  | 
|
| 52059 | 502  | 
if Config.get lthy quick_and_dirty then  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
503  | 
replicate (n + live) (K all_tac)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
504  | 
else  | 
| 49304 | 505  | 
replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
506  | 
(map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
507  | 
|
| 49303 | 508  | 
val in_alt_thm =  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
509  | 
let  | 
| 49303 | 510  | 
val inx = mk_in Asets sets T;  | 
511  | 
val in_alt = mk_in (drop n Asets) bnf_sets T;  | 
|
512  | 
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));  | 
|
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
513  | 
in  | 
| 51551 | 514  | 
Goal.prove_sorry lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
515  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
516  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
517  | 
fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
518  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
519  | 
fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm;  | 
| 49456 | 520  | 
|
| 53289 | 521  | 
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
522  | 
bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
523  | 
|
| 49303 | 524  | 
val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
525  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
526  | 
fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
527  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
528  | 
val (bnf', lthy') =  | 
| 
51767
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
529  | 
bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty  | 
| 54421 | 530  | 
[] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
531  | 
in  | 
| 49503 | 532  | 
(bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
533  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
534  | 
|
| 55904 | 535  | 
fun lift_bnf qualify n bnf (accum as ((cache, unfold_set), lthy)) =  | 
536  | 
let val key = key_of_lift n bnf in  | 
|
537  | 
(case Typtab.lookup cache key of  | 
|
538  | 
SOME (bnf, _) => (bnf, accum)  | 
|
539  | 
| NONE => cache_comp_simple key cache (raw_lift_bnf qualify n bnf (unfold_set, lthy)))  | 
|
540  | 
end;  | 
|
541  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
542  | 
(* Changing the order of live variables *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
543  | 
|
| 55904 | 544  | 
fun raw_permute_bnf qualify src dest bnf (accum as (unfold_set, lthy)) =  | 
545  | 
if src = dest then (bnf, accum) else  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
546  | 
let  | 
| 49425 | 547  | 
val b = Binding.suffix_name (mk_permuteN src dest) (name_of_bnf bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
548  | 
val live = live_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
549  | 
val dead = dead_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
550  | 
val nwits = nwits_of_bnf bnf;  | 
| 
55480
 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 
blanchet 
parents: 
55197 
diff
changeset
 | 
551  | 
|
| 
 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 
blanchet 
parents: 
55197 
diff
changeset
 | 
552  | 
fun permute xs = permute_like_unique (op =) src dest xs;  | 
| 
 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 
blanchet 
parents: 
55197 
diff
changeset
 | 
553  | 
fun unpermute xs = permute_like_unique (op =) dest src xs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
554  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
555  | 
val (Ds, lthy1) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
556  | 
(Variable.invent_types (replicate dead HOLogic.typeS) lthy);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
557  | 
val (As, lthy2) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
558  | 
(Variable.invent_types (replicate live HOLogic.typeS) lthy1);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
559  | 
val (Bs, _(*lthy3*)) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
560  | 
(Variable.invent_types (replicate live HOLogic.typeS) lthy2);  | 
| 
 
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  | 
val (Asets, _(*names_lthy*)) = lthy  | 
| 49456 | 563  | 
|> mk_Frees "A" (map HOLogic.mk_setT (permute As));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
564  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
565  | 
val T = mk_T_of_bnf Ds As bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
566  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
567  | 
(*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*)  | 
| 49303 | 568  | 
val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs))  | 
| 53038 | 569  | 
(Term.list_comb (mk_map_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0))));  | 
| 49507 | 570  | 
(*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*)  | 
571  | 
val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs))  | 
|
| 53038 | 572  | 
(Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0))));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
573  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
574  | 
val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;  | 
| 49303 | 575  | 
val sets = permute bnf_sets;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
576  | 
|
| 49303 | 577  | 
val bd = mk_bd_of_bnf Ds As bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
578  | 
|
| 53270 | 579  | 
fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;  | 
| 53287 | 580  | 
fun map_comp0_tac _ = rtac (map_comp0_of_bnf bnf) 1;  | 
| 55197 | 581  | 
fun map_cong0_tac ctxt =  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
582  | 
rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);  | 
| 53289 | 583  | 
val set_map0_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf));  | 
| 49303 | 584  | 
fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;  | 
585  | 
fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;  | 
|
586  | 
val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
587  | 
|
| 49303 | 588  | 
val in_alt_thm =  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
589  | 
let  | 
| 49303 | 590  | 
val inx = mk_in Asets sets T;  | 
| 53038 | 591  | 
val in_alt = mk_in (unpermute Asets) bnf_sets T;  | 
| 49303 | 592  | 
val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
593  | 
in  | 
| 51551 | 594  | 
Goal.prove_sorry lthy [] [] goal (K (mk_permute_in_alt_tac src dest))  | 
| 
49155
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
595  | 
|> Thm.close_derivation  | 
| 
 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 
traytel 
parents: 
49123 
diff
changeset
 | 
596  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
597  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
598  | 
fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
599  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
600  | 
fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm;  | 
| 49456 | 601  | 
|
| 53289 | 602  | 
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54742 
diff
changeset
 | 
603  | 
bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
604  | 
|
| 49303 | 605  | 
val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
606  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
607  | 
fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
608  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
609  | 
val (bnf', lthy') =  | 
| 
51767
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
610  | 
bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty  | 
| 54421 | 611  | 
[] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
612  | 
in  | 
| 49503 | 613  | 
(bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
614  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
615  | 
|
| 55904 | 616  | 
fun permute_bnf qualify src dest bnf (accum as ((cache, unfold_set), lthy)) =  | 
617  | 
let val key = key_of_permute src dest bnf in  | 
|
618  | 
(case Typtab.lookup cache key of  | 
|
619  | 
SOME (bnf, _) => (bnf, accum)  | 
|
620  | 
| NONE => cache_comp_simple key cache (raw_permute_bnf qualify src dest bnf (unfold_set, lthy)))  | 
|
621  | 
end;  | 
|
622  | 
||
| 49014 | 623  | 
(* Composition pipeline *)  | 
624  | 
||
625  | 
fun permute_and_kill qualify n src dest bnf =  | 
|
| 
55703
 
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
 
blanchet 
parents: 
55480 
diff
changeset
 | 
626  | 
permute_bnf qualify src dest bnf  | 
| 49304 | 627  | 
#> uncurry (kill_bnf qualify n);  | 
| 49014 | 628  | 
|
629  | 
fun lift_and_permute qualify n src dest bnf =  | 
|
| 
55703
 
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
 
blanchet 
parents: 
55480 
diff
changeset
 | 
630  | 
lift_bnf qualify n bnf  | 
| 49014 | 631  | 
#> uncurry (permute_bnf qualify src dest);  | 
632  | 
||
| 55904 | 633  | 
fun normalize_bnfs qualify Ass Ds sort bnfs accum =  | 
| 49014 | 634  | 
let  | 
635  | 
val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass;  | 
|
| 52985 | 636  | 
val kill_poss = map (find_indices op = Ds) Ass;  | 
637  | 
val live_poss = map2 (subtract op =) kill_poss before_kill_src;  | 
|
| 49014 | 638  | 
val before_kill_dest = map2 append kill_poss live_poss;  | 
639  | 
val kill_ns = map length kill_poss;  | 
|
| 55904 | 640  | 
val (inners', accum') =  | 
| 49014 | 641  | 
fold_map5 (fn i => permute_and_kill (qualify i))  | 
642  | 
(if length bnfs = 1 then [0] else (1 upto length bnfs))  | 
|
| 55904 | 643  | 
kill_ns before_kill_src before_kill_dest bnfs accum;  | 
| 49014 | 644  | 
|
645  | 
val Ass' = map2 (map o nth) Ass live_poss;  | 
|
646  | 
val As = sort Ass';  | 
|
647  | 
val after_lift_dest = replicate (length Ass') (0 upto (length As - 1));  | 
|
648  | 
val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass';  | 
|
| 52985 | 649  | 
val new_poss = map2 (subtract op =) old_poss after_lift_dest;  | 
| 49014 | 650  | 
val after_lift_src = map2 append new_poss old_poss;  | 
651  | 
val lift_ns = map (fn xs => length As - length xs) Ass';  | 
|
652  | 
in  | 
|
653  | 
((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i))  | 
|
| 
55703
 
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
 
blanchet 
parents: 
55480 
diff
changeset
 | 
654  | 
(if length bnfs = 1 then [0] else 1 upto length bnfs)  | 
| 55904 | 655  | 
lift_ns after_lift_src after_lift_dest inners' accum')  | 
| 49014 | 656  | 
end;  | 
657  | 
||
658  | 
fun default_comp_sort Ass =  | 
|
659  | 
Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []);  | 
|
660  | 
||
| 55904 | 661  | 
fun raw_compose_bnf const_policy qualify sort outer inners oDs Dss tfreess accum =  | 
| 49014 | 662  | 
let  | 
| 49425 | 663  | 
val b = name_of_bnf outer;  | 
| 49014 | 664  | 
|
| 49121 | 665  | 
val Ass = map (map Term.dest_TFree) tfreess;  | 
| 49014 | 666  | 
val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];  | 
667  | 
||
| 55904 | 668  | 
val ((kill_poss, As), (inners', ((cache', unfold_set'), lthy'))) =  | 
669  | 
normalize_bnfs qualify Ass Ds sort inners accum;  | 
|
| 49014 | 670  | 
|
671  | 
val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss);  | 
|
672  | 
val As = map TFree As;  | 
|
673  | 
in  | 
|
| 49425 | 674  | 
apfst (rpair (Ds, As))  | 
| 55904 | 675  | 
(apsnd (apfst (pair cache'))  | 
676  | 
(clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy')))  | 
|
677  | 
end;  | 
|
678  | 
||
679  | 
fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (accum as ((cache, _), _)) =  | 
|
680  | 
let val key = key_of_compose oDs Dss tfreess outer inners in  | 
|
681  | 
(case Typtab.lookup cache key of  | 
|
682  | 
SOME bnf_Ds_As => (bnf_Ds_As, accum)  | 
|
683  | 
| NONE =>  | 
|
684  | 
cache_comp key (raw_compose_bnf const_policy qualify sort outer inners oDs Dss tfreess accum))  | 
|
| 49014 | 685  | 
end;  | 
686  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
687  | 
(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
688  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
689  | 
type absT_info =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
690  | 
  {absT: typ,
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
691  | 
repT: typ,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
692  | 
abs: term,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
693  | 
rep: term,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
694  | 
abs_inject: thm,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
695  | 
abs_inverse: thm,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
696  | 
type_definition: thm};  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
697  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
698  | 
fun morph_absT_info phi  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
699  | 
  {absT, repT, abs, rep, abs_inject, abs_inverse, type_definition} =
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
700  | 
  {absT = Morphism.typ phi absT,
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
701  | 
repT = Morphism.typ phi repT,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
702  | 
abs = Morphism.term phi abs,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
703  | 
rep = Morphism.term phi rep,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
704  | 
abs_inject = Morphism.thm phi abs_inject,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
705  | 
abs_inverse = Morphism.thm phi abs_inverse,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
706  | 
type_definition = Morphism.thm phi type_definition};  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
707  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
708  | 
fun mk_absT thy repT absT repU =  | 
| 
55900
 
21aa30ea6806
propagate the exception that is expected later on
 
traytel 
parents: 
55856 
diff
changeset
 | 
709  | 
let  | 
| 
 
21aa30ea6806
propagate the exception that is expected later on
 
traytel 
parents: 
55856 
diff
changeset
 | 
710  | 
val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (repT, repU) Vartab.empty) [];  | 
| 
 
21aa30ea6806
propagate the exception that is expected later on
 
traytel 
parents: 
55856 
diff
changeset
 | 
711  | 
in Term.typ_subst_TVars rho absT end  | 
| 
 
21aa30ea6806
propagate the exception that is expected later on
 
traytel 
parents: 
55856 
diff
changeset
 | 
712  | 
  handle Type.TYPE_MATCH => raise Term.TYPE ("mk_absT", [repT, absT, repU], []);
 | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
713  | 
|
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
714  | 
fun mk_repT absT repT absU =  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
715  | 
if absT = repT then absU  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
716  | 
else  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
717  | 
(case (absT, absU) of  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
718  | 
(Type (C, Ts), Type (C', Us)) =>  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
719  | 
if C = C' then Term.typ_subst_atomic (Ts ~~ Us) repT  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
720  | 
        else raise Term.TYPE ("mk_repT", [absT, repT, absT], [])
 | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
721  | 
    | _ => raise Term.TYPE ("mk_repT", [absT, repT, absT], []));
 | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
722  | 
|
| 
55855
 
98ad5680173a
use same identity function for abs and rep (doesn't seem to confuse any proofs)
 
blanchet 
parents: 
55854 
diff
changeset
 | 
723  | 
fun mk_abs_or_rep _ absU (Const (@{const_name id_bnf_comp}, _)) =
 | 
| 
 
98ad5680173a
use same identity function for abs and rep (doesn't seem to confuse any proofs)
 
blanchet 
parents: 
55854 
diff
changeset
 | 
724  | 
    Const (@{const_name id_bnf_comp}, absU --> absU)
 | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
725  | 
| mk_abs_or_rep getT (Type (_, Us)) abs =  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
726  | 
let val Ts = snd (dest_Type (getT (fastype_of abs)))  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
727  | 
in Term.subst_atomic_types (Ts ~~ Us) abs end;  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
728  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
729  | 
val mk_abs = mk_abs_or_rep range_type;  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
730  | 
val mk_rep = mk_abs_or_rep domain_type;  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
731  | 
|
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
732  | 
val smart_max_inline_type_size = 5; (*FUDGE*)  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
733  | 
|
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
734  | 
fun maybe_typedef (b, As, mx) set opt_morphs tac =  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
735  | 
let  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
736  | 
val repT = HOLogic.dest_setT (fastype_of set);  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
737  | 
val inline = Term.size_of_typ repT <= smart_max_inline_type_size;  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
738  | 
in  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
739  | 
if inline then  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
740  | 
pair (repT,  | 
| 
55855
 
98ad5680173a
use same identity function for abs and rep (doesn't seem to confuse any proofs)
 
blanchet 
parents: 
55854 
diff
changeset
 | 
741  | 
        (@{const_name id_bnf_comp}, @{const_name id_bnf_comp},
 | 
| 
 
98ad5680173a
use same identity function for abs and rep (doesn't seem to confuse any proofs)
 
blanchet 
parents: 
55854 
diff
changeset
 | 
742  | 
         @{thm type_definition_id_bnf_comp_UNIV},
 | 
| 
 
98ad5680173a
use same identity function for abs and rep (doesn't seem to confuse any proofs)
 
blanchet 
parents: 
55854 
diff
changeset
 | 
743  | 
         @{thm type_definition.Abs_inverse[OF type_definition_id_bnf_comp_UNIV]},
 | 
| 
 
98ad5680173a
use same identity function for abs and rep (doesn't seem to confuse any proofs)
 
blanchet 
parents: 
55854 
diff
changeset
 | 
744  | 
         @{thm type_definition.Abs_inject[OF type_definition_id_bnf_comp_UNIV]}))
 | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
745  | 
else  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
746  | 
typedef (b, As, mx) set opt_morphs tac  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
747  | 
      #>> (fn (T_name, ({Rep_name, Abs_name, ...},
 | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
748  | 
          {type_definition, Abs_inverse, Abs_inject, ...}) : Typedef.info) =>
 | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
749  | 
(Type (T_name, map TFree As), (Rep_name, Abs_name, type_definition, Abs_inverse, Abs_inject)))  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
750  | 
end;  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
751  | 
|
| 53264 | 752  | 
fun seal_bnf qualify (unfold_set : unfold_set) b Ds bnf lthy =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
753  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
754  | 
val live = live_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
755  | 
val nwits = nwits_of_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
756  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
757  | 
val (As, lthy1) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
758  | 
(Variable.invent_types (replicate live HOLogic.typeS) (fold Variable.declare_typ Ds lthy));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
759  | 
val (Bs, _) = apfst (map TFree)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
760  | 
(Variable.invent_types (replicate live HOLogic.typeS) lthy1);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
761  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
762  | 
val (((fs, fs'), (Rs, Rs')), _(*names_lthy*)) = lthy  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
763  | 
|> mk_Frees' "f" (map2 (curry op -->) As Bs)  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
764  | 
||>> mk_Frees' "R" (map2 mk_pred2T As Bs)  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
765  | 
|
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
766  | 
val expand_id_bnf_comp_def =  | 
| 
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
767  | 
expand_term_const [Thm.prop_of id_bnf_comp_def |> Logic.dest_equals]  | 
| 
49713
 
3d07ddf70f8b
do not expose details of internal data structures for composition of BNFs
 
traytel 
parents: 
49669 
diff
changeset
 | 
768  | 
val map_unfolds = #map_unfolds unfold_set;  | 
| 
 
3d07ddf70f8b
do not expose details of internal data structures for composition of BNFs
 
traytel 
parents: 
49669 
diff
changeset
 | 
769  | 
val set_unfoldss = #set_unfoldss unfold_set;  | 
| 
 
3d07ddf70f8b
do not expose details of internal data structures for composition of BNFs
 
traytel 
parents: 
49669 
diff
changeset
 | 
770  | 
val rel_unfolds = #rel_unfolds unfold_set;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
771  | 
|
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
772  | 
val expand_maps = expand_id_bnf_comp_def o  | 
| 49507 | 773  | 
fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds);  | 
774  | 
val expand_sets =  | 
|
775  | 
fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss);  | 
|
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
776  | 
val expand_rels = expand_id_bnf_comp_def o  | 
| 49507 | 777  | 
fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) rel_unfolds);  | 
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
778  | 
fun unfold_maps ctxt = fold (unfold_thms ctxt o single) (id_bnf_comp_def :: map_unfolds);  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54421 
diff
changeset
 | 
779  | 
fun unfold_sets ctxt = fold (unfold_thms ctxt) set_unfoldss;  | 
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
780  | 
fun unfold_rels ctxt = fold (unfold_thms ctxt o single) (id_bnf_comp_def :: rel_unfolds);  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54421 
diff
changeset
 | 
781  | 
fun unfold_all ctxt = unfold_sets ctxt o unfold_maps ctxt o unfold_rels ctxt;  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
782  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
783  | 
val repTA = mk_T_of_bnf Ds As bnf;  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
784  | 
val T_bind = qualify b;  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
785  | 
val TA_params = Term.add_tfreesT repTA [];  | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
786  | 
val ((TA, (Rep_name, Abs_name, type_definition, Abs_inverse, Abs_inject)), lthy) =  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
787  | 
maybe_typedef (T_bind, TA_params, NoSyn)  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
788  | 
(HOLogic.mk_UNIV repTA) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;  | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
789  | 
|
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
790  | 
val repTB = mk_T_of_bnf Ds Bs bnf;  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
791  | 
val TB = Term.typ_subst_atomic (As ~~ Bs) TA;  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
792  | 
val RepA = Const (Rep_name, TA --> repTA);  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
793  | 
val RepB = Const (Rep_name, TB --> repTB);  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
794  | 
val AbsA = Const (Abs_name, repTA --> TA);  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
795  | 
val AbsB = Const (Abs_name, repTB --> TB);  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
796  | 
    val Abs_inject' = Abs_inject OF @{thms UNIV_I UNIV_I};
 | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
797  | 
    val Abs_inverse' = Abs_inverse OF @{thms UNIV_I};
 | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
798  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
799  | 
    val absT_info = {absT = TA, repT = repTA, abs = AbsA, rep = RepA, abs_inject = Abs_inject',
 | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
800  | 
abs_inverse = Abs_inverse', type_definition = type_definition};  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
801  | 
|
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
802  | 
val bnf_map = fold_rev Term.absfree fs' (HOLogic.mk_comp (HOLogic.mk_comp (AbsB,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
803  | 
Term.list_comb (expand_maps (mk_map_of_bnf Ds As Bs bnf), fs)), RepA));  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
804  | 
val bnf_sets = map ((fn t => HOLogic.mk_comp (t, RepA)) o expand_maps o expand_sets)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
805  | 
(mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
806  | 
val bnf_bd = mk_bd_of_bnf Ds As bnf;  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
807  | 
val bnf_rel = fold_rev Term.absfree Rs' (mk_vimage2p RepA RepB $  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
808  | 
(Term.list_comb (expand_rels (mk_rel_of_bnf Ds As Bs bnf), Rs)));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
809  | 
|
| 
55704
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
810  | 
(*bd may depend only on dead type variables*)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
811  | 
val bd_repT = fst (dest_relT (fastype_of bnf_bd));  | 
| 53264 | 812  | 
    val bdT_bind = qualify (Binding.suffix_name ("_" ^ bdTN) b);
 | 
| 
55707
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
813  | 
val params = Term.add_tfreesT bd_repT [];  | 
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49155 
diff
changeset
 | 
814  | 
val deads = map TFree params;  | 
| 
55707
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
815  | 
val all_deads = map TFree (fold Term.add_tfreesT Ds []);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
816  | 
|
| 49228 | 817  | 
val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) =  | 
| 49835 | 818  | 
typedef (bdT_bind, params, NoSyn)  | 
| 49228 | 819  | 
(HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
820  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
821  | 
val bnf_bd' = mk_dir_image bnf_bd  | 
| 
49185
 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 
traytel 
parents: 
49155 
diff
changeset
 | 
822  | 
(Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, deads)))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
823  | 
|
| 49228 | 824  | 
val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info);  | 
825  | 
val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj (#Abs_cases bdT_loc_info);  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
826  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
827  | 
    val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf];
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
828  | 
val bd_card_order =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
829  | 
      @{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf];
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
830  | 
val bd_cinfinite =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
831  | 
      (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
832  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
833  | 
fun map_id0_tac ctxt =  | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
834  | 
      rtac (@{thm type_copy_map_id0} OF [type_definition, unfold_maps ctxt (map_id0_of_bnf bnf)]) 1;
 | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
835  | 
fun map_comp0_tac ctxt =  | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
836  | 
      rtac (@{thm type_copy_map_comp0} OF
 | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
837  | 
[type_definition, unfold_maps ctxt (map_comp0_of_bnf bnf)]) 1;  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
838  | 
fun map_cong0_tac ctxt =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
839  | 
      EVERY' (rtac @{thm type_copy_map_cong0} :: rtac (unfold_all ctxt (map_cong0_of_bnf bnf)) ::
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
840  | 
map (fn i => EVERY' [select_prem_tac live (dtac meta_spec) i, etac meta_mp,  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
841  | 
etac (o_apply RS equalityD2 RS set_mp)]) (1 upto live)) 1;  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
842  | 
fun set_map0_tac thm ctxt =  | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
843  | 
      rtac (@{thm type_copy_set_map0} OF [type_definition, unfold_all ctxt thm]) 1;
 | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
844  | 
    val set_bd_tacs = map (fn thm => fn ctxt => rtac (@{thm ordLeq_ordIso_trans} OF
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
845  | 
        [unfold_sets ctxt thm, bd_ordIso] RS @{thm type_copy_set_bd}) 1)
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
846  | 
(set_bd_of_bnf bnf);  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
847  | 
fun le_rel_OO_tac ctxt =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
848  | 
      rtac (unfold_rels ctxt (le_rel_OO_of_bnf bnf) RS @{thm vimage2p_relcompp_mono}) 1;
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
849  | 
fun rel_OO_Grp_tac ctxt =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
850  | 
      (rtac (unfold_all ctxt (rel_OO_Grp_of_bnf bnf) RS @{thm vimage2p_cong} RS trans) THEN'
 | 
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
851  | 
SELECT_GOAL (unfold_thms_tac ctxt [o_apply,  | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
852  | 
        type_definition RS @{thm type_copy_vimage2p_Grp_Rep},
 | 
| 
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
853  | 
        type_definition RS @{thm vimage2p_relcompp_converse}]) THEN' rtac refl) 1;
 | 
| 49456 | 854  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
855  | 
val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
856  | 
(map set_map0_tac (set_map0_of_bnf bnf)) (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1))  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
857  | 
set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
858  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
859  | 
val bnf_wits = map (fn (I, t) =>  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
860  | 
fold Term.absdummy (map (nth As) I)  | 
| 
55935
 
8f20d09d294e
move special BNFs used for composition only to BNF_Comp;
 
traytel 
parents: 
55930 
diff
changeset
 | 
861  | 
(AbsA $ Term.list_comb (expand_id_bnf_comp_def t, map Bound (0 upto length I - 1))))  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
862  | 
(mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
863  | 
|
| 
55854
 
ee270328a781
make 'typedef' optional, depending on size of original type
 
blanchet 
parents: 
55853 
diff
changeset
 | 
864  | 
    fun wit_tac ctxt = ALLGOALS (dtac (type_definition RS @{thm type_copy_wit})) THEN
 | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54421 
diff
changeset
 | 
865  | 
mk_simple_wit_tac (map (unfold_all ctxt) (wit_thms_of_bnf bnf));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
866  | 
|
| 53264 | 867  | 
val (bnf', lthy') =  | 
| 
55707
 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 
traytel 
parents: 
55706 
diff
changeset
 | 
868  | 
bnf_def Hardly_Inline (user_policy Dont_Note) qualify tacs wit_tac (SOME all_deads)  | 
| 53264 | 869  | 
Binding.empty Binding.empty []  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
870  | 
((((((b, TA), bnf_map), bnf_sets), bnf_bd'), bnf_wits), SOME bnf_rel) lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
871  | 
in  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55707 
diff
changeset
 | 
872  | 
((bnf', (all_deads, absT_info)), lthy')  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
873  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
874  | 
|
| 53222 | 875  | 
exception BAD_DEAD of typ * typ;  | 
876  | 
||
| 
55704
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
877  | 
fun bnf_of_typ _ _ _ _ Ds0 (T as TFree T') accum =  | 
| 
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
878  | 
(if member (op =) Ds0 T' then (DEADID_bnf, ([T], [])) else (ID_bnf, ([], [T])), accum)  | 
| 
55703
 
a21069381ba5
optimization of 'bnf_of_typ' if all variables are dead
 
blanchet 
parents: 
55480 
diff
changeset
 | 
879  | 
| bnf_of_typ _ _ _ _ _ (TVar _) _ = error "Unexpected schematic variable"  | 
| 55706 | 880  | 
| bnf_of_typ const_policy qualify' sort Xs Ds0 (T as Type (C, Ts)) (accum as (_, lthy)) =  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
881  | 
let  | 
| 53222 | 882  | 
fun check_bad_dead ((_, (deads, _)), _) =  | 
883  | 
let val Ds = fold Term.add_tfreesT deads [] in  | 
|
884  | 
(case Library.inter (op =) Ds Xs of [] => ()  | 
|
| 55705 | 885  | 
| X :: _ => raise BAD_DEAD (TFree X, T))  | 
| 53222 | 886  | 
end;  | 
887  | 
||
| 
55704
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
888  | 
val tfrees = subtract (op =) Ds0 (Term.add_tfreesT T []);  | 
| 
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
889  | 
val bnf_opt = if null tfrees then NONE else bnf_of lthy C;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
890  | 
in  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
891  | 
(case bnf_opt of  | 
| 
55704
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
892  | 
NONE => ((DEADID_bnf, ([T], [])), accum)  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
893  | 
| SOME bnf =>  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
894  | 
if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
895  | 
let  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
896  | 
val T' = T_of_bnf bnf;  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
897  | 
val deads = deads_of_bnf bnf;  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
898  | 
val lives = lives_of_bnf bnf;  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
899  | 
val tvars' = Term.add_tvarsT T' [];  | 
| 55904 | 900  | 
val Ds_As =  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
901  | 
pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees)))  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
902  | 
(deads, lives);  | 
| 55904 | 903  | 
in ((bnf, Ds_As), accum) end  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
904  | 
else  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
905  | 
let  | 
| 49425 | 906  | 
val name = Long_Name.base_name C;  | 
907  | 
fun qualify i =  | 
|
908  | 
let val namei = name ^ nonzero_string_of_int i;  | 
|
909  | 
in qualify' o Binding.qualify true namei end;  | 
|
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
910  | 
val odead = dead_of_bnf bnf;  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
911  | 
val olive = live_of_bnf bnf;  | 
| 52985 | 912  | 
            val oDs_pos = find_indices op = [TFree ("dead", [])] (snd (Term.dest_Type
 | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
913  | 
              (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
 | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
914  | 
val oDs = map (nth Ts) oDs_pos;  | 
| 
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
915  | 
val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1));  | 
| 55904 | 916  | 
val ((inners, (Dss, Ass)), (accum', lthy')) =  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
917  | 
apfst (apsnd split_list o split_list)  | 
| 
55704
 
a97b9b81e195
optimized 'bnf_of_typ' further w.r.t. dead variables
 
blanchet 
parents: 
55703 
diff
changeset
 | 
918  | 
(fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort Xs Ds0)  | 
| 55706 | 919  | 
(if length Ts' = 1 then [0] else (1 upto length Ts')) Ts' accum);  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
920  | 
in  | 
| 55904 | 921  | 
compose_bnf const_policy qualify sort bnf inners oDs Dss Ass (accum', lthy')  | 
| 
49186
 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 
traytel 
parents: 
49185 
diff
changeset
 | 
922  | 
end)  | 
| 53222 | 923  | 
|> tap check_bad_dead  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
924  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
925  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
926  | 
end;  |