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