| author | sultana | 
| Tue, 03 Sep 2013 21:46:40 +0100 | |
| changeset 53385 | 7edd43d0c0ba | 
| parent 53289 | 5e0623448bdb | 
| child 54421 | 632be352a5a3 | 
| permissions | -rw-r--r-- | 
| 49509 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 blanchet parents: 
49507diff
changeset | 1 | (* Title: HOL/BNF/Tools/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 | 
| 51837 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 blanchet parents: 
51782diff
changeset | 11 | 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 | 12 | val DEADID_bnf: BNF_Def.bnf | 
| 49585 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49538diff
changeset | 13 | |
| 49502 | 14 | type unfold_set | 
| 15 | 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 | 16 | |
| 53222 | 17 | exception BAD_DEAD of typ * typ | 
| 18 | ||
| 49425 | 19 | val bnf_of_typ: BNF_Def.const_policy -> (binding -> binding) -> | 
| 53222 | 20 | ((string * sort) list list -> (string * sort) list) -> (string * sort) list -> typ -> | 
| 21 | unfold_set * Proof.context -> | |
| 51837 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 blanchet parents: 
51782diff
changeset | 22 | (BNF_Def.bnf * (typ list * typ list)) * (unfold_set * Proof.context) | 
| 49014 | 23 | val default_comp_sort: (string * sort) list list -> (string * sort) list | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 24 | val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list -> | 
| 51837 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 blanchet parents: 
51782diff
changeset | 25 |     (''a list list -> ''a list) -> BNF_Def.bnf list -> unfold_set -> Proof.context ->
 | 
| 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 blanchet parents: 
51782diff
changeset | 26 | (int list list * ''a list) * (BNF_Def.bnf list * (unfold_set * Proof.context)) | 
| 53264 | 27 | val seal_bnf: (binding -> binding) -> unfold_set -> binding -> typ list -> BNF_Def.bnf -> | 
| 28 | Proof.context -> (BNF_Def.bnf * typ list) * local_theory | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 29 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 30 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 31 | structure BNF_Comp : BNF_COMP = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 32 | struct | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 33 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 34 | open BNF_Def | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 35 | open BNF_Util | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 36 | open BNF_Tactics | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 37 | open BNF_Comp_Tactics | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 38 | |
| 49585 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49538diff
changeset | 39 | val ID_bnf = the (bnf_of @{context} "Basic_BNFs.ID");
 | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49538diff
changeset | 40 | val DEADID_bnf = the (bnf_of @{context} "Basic_BNFs.DEADID");
 | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49538diff
changeset | 41 | |
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49538diff
changeset | 42 | (* TODO: Replace by "BNF_Defs.defs list" *) | 
| 49502 | 43 | type unfold_set = {
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 44 | map_unfolds: thm list, | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 45 | set_unfoldss: thm list list, | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 46 | rel_unfolds: thm list | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 47 | }; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 48 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 49 | 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 | 50 | |
| 49503 | 51 | fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new; | 
| 52 | fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms; | |
| 53 | ||
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 54 | 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 | 55 |   {map_unfolds, set_unfoldss, rel_unfolds} =
 | 
| 49503 | 56 |   {map_unfolds = add_to_thms map_unfolds map,
 | 
| 57 | 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 | 58 | 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 | 59 | |
| 49503 | 60 | 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 | 61 | 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 | 62 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 63 | val bdTN = "bdT"; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 64 | |
| 49425 | 65 | fun mk_killN n = "_kill" ^ string_of_int n; | 
| 66 | 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 | 67 | fun mk_permuteN src dest = | 
| 49425 | 68 | "_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 | 69 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 70 | (*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 | 71 | fun expand_term_const defs = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 72 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 73 | 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 | 74 | 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 | 75 | 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 | 76 | |
| 49502 | 77 | 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 | 78 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 79 | 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 | 80 | val onwits = nwits_of_bnf outer; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 81 | val odead = dead_of_bnf outer; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 82 | val inner = hd inners; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 83 | val ilive = live_of_bnf inner; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 84 | val ideads = map dead_of_bnf inners; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 85 | 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 | 86 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 87 | (* 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 | 88 | 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 | 89 | 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 | 90 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 91 | val (oDs, lthy1) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 92 | (Variable.invent_types (replicate odead HOLogic.typeS) lthy); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 93 | val (Dss, lthy2) = apfst (map (map TFree)) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 94 | (fold_map Variable.invent_types (map (fn n => replicate n HOLogic.typeS) ideads) lthy1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 95 | val (Ass, lthy3) = apfst (replicate ilive o map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 96 | (Variable.invent_types (replicate ilive HOLogic.typeS) lthy2); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 97 | 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 | 98 | val Ass_repl = replicate olive As; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 99 | val (Bs, _(*lthy4*)) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 100 | (Variable.invent_types (replicate ilive HOLogic.typeS) lthy3); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 101 | 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 | 102 | |
| 49463 | 103 | val ((((fs', Qs'), Asets), xs), _(*names_lthy*)) = lthy | 
| 52923 | 104 | |> apfst snd o mk_Frees' "f" (map2 (curry op -->) As Bs) | 
| 49463 | 105 | ||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs) | 
| 49456 | 106 | ||>> 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 | 107 | ||>> mk_Frees "x" As; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 108 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 109 | val CAs = map3 mk_T_of_bnf Dss Ass_repl inners; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 110 | val CCA = mk_T_of_bnf oDs CAs outer; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 111 | val CBs = map3 mk_T_of_bnf Dss Bss_repl inners; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 112 | val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 113 | val inner_setss = map3 mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 114 | val inner_bds = map3 mk_bd_of_bnf Dss Ass_repl inners; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 115 | 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 | 116 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 117 | (*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*) | 
| 49303 | 118 | 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 | 119 | (Term.list_comb (mk_map_of_bnf oDs CAs CBs outer, | 
| 49463 | 120 | 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 | 121 | mk_map_of_bnf Ds As Bs) Dss inners)); | 
| 49507 | 122 | (*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*) | 
| 123 | val rel = fold_rev Term.abs Qs' | |
| 124 | (Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer, | |
| 49463 | 125 | map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o | 
| 49507 | 126 | mk_rel_of_bnf Ds As Bs) Dss inners)); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 127 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 128 |     (*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 | 129 |     (*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*)
 | 
| 49303 | 130 | fun mk_set i = | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 131 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 132 | 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 | 133 | val outer_set = mk_collect | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 134 | (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 | 135 | (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 | 136 | 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 | 137 | 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 | 138 | 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 | 139 | val collect_image = mk_collect | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 140 | (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 | 141 | (CCA --> HOLogic.mk_setT T); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 142 | in | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 143 | (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 | 144 | 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 | 145 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 146 | |
| 49303 | 147 | 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 | 148 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 149 | (*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*) | 
| 49303 | 150 | val bd = Term.absdummy CCA (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 | 151 | |
| 53270 | 152 | fun map_id0_tac _ = | 
| 153 | mk_comp_map_id0_tac (map_id0_of_bnf outer) (map_cong0_of_bnf outer) | |
| 154 | (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 | 155 | |
| 53287 | 156 | fun map_comp0_tac _ = | 
| 157 | mk_comp_map_comp0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer) | |
| 158 | (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 | 159 | |
| 53289 | 160 | fun mk_single_set_map0_tac i _ = | 
| 161 | mk_comp_set_map0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer) | |
| 51766 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 blanchet parents: 
51761diff
changeset | 162 | (collect_set_map_of_bnf outer) | 
| 53289 | 163 | (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 | 164 | |
| 53289 | 165 | 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 | 166 | |
| 49303 | 167 | fun bd_card_order_tac _ = | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 168 | mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 169 | |
| 49303 | 170 | fun bd_cinfinite_tac _ = | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 171 | mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 172 | |
| 49303 | 173 | val set_alt_thms = | 
| 52059 | 174 | if Config.get lthy quick_and_dirty then | 
| 49456 | 175 | [] | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 176 | else | 
| 49109 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 traytel parents: 
49019diff
changeset | 177 | map (fn goal => | 
| 51551 | 178 | Goal.prove_sorry lthy [] [] goal | 
| 49714 | 179 |             (fn {context = ctxt, prems = _} =>
 | 
| 51766 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 blanchet parents: 
51761diff
changeset | 180 | 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 | 181 | |> Thm.close_derivation) | 
| 49303 | 182 | (map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) sets sets_alt); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 183 | |
| 51761 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 184 | fun map_cong0_tac _ = | 
| 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 185 | mk_comp_map_cong0_tac set_alt_thms (map_cong0_of_bnf outer) (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 | 186 | |
| 49303 | 187 | val set_bd_tacs = | 
| 52059 | 188 | if Config.get lthy quick_and_dirty then | 
| 49669 | 189 | 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 | 190 | else | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 191 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 192 | 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 | 193 | 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 | 194 | val inner_bd_Card_orders = map bd_Card_order_of_bnf inners; | 
| 49303 | 195 | 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 | 196 |             @{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 | 197 | nth outer_set_bds j] | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 198 | val single_set_bd_thmss = | 
| 49303 | 199 | 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 | 200 | in | 
| 49714 | 201 |           map2 (fn set_alt => fn single_set_bds => fn {context = ctxt, prems = _} =>
 | 
| 202 | mk_comp_set_bd_tac ctxt set_alt single_set_bds) | |
| 49303 | 203 | 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 | 204 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 205 | |
| 49303 | 206 | val in_alt_thm = | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 207 | let | 
| 49303 | 208 | val inx = mk_in Asets sets CCA; | 
| 209 | val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA; | |
| 210 | 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 | 211 | in | 
| 51551 | 212 | Goal.prove_sorry lthy [] [] goal | 
| 49714 | 213 |           (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 | 214 | |> Thm.close_derivation | 
| 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 215 | end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 216 | |
| 49303 | 217 | fun map_wpull_tac _ = | 
| 218 | mk_map_wpull_tac in_alt_thm (map map_wpull_of_bnf inners) (map_wpull_of_bnf outer); | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 219 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 220 | fun rel_OO_Grp_tac _ = | 
| 49456 | 221 | let | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 222 | val outer_rel_Grp = rel_Grp_of_bnf outer RS sym; | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 223 | val outer_rel_cong = rel_cong_of_bnf outer; | 
| 49463 | 224 | val thm = | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 225 |           (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 | 226 |              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 | 227 |                [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 | 228 | rel_conversep_of_bnf outer RS sym], outer_rel_Grp], | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 229 | trans OF [rel_OO_of_bnf outer RS sym, outer_rel_cong OF | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 230 | (map (fn bnf => rel_OO_Grp_of_bnf bnf RS sym) inners)]]] RS sym) | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 231 | (*|> unfold_thms lthy (rel_def_of_bnf outer :: map rel_def_of_bnf inners)*); | 
| 49456 | 232 | in | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 233 | rtac thm 1 | 
| 49463 | 234 | end; | 
| 49456 | 235 | |
| 53289 | 236 | val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52059diff
changeset | 237 | bd_cinfinite_tac set_bd_tacs map_wpull_tac rel_OO_Grp_tac; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 238 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 239 | 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 | 240 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 241 | val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I))) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 242 | (map3 (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As)) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 243 | Dss inwitss inners); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 244 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 245 | val 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 | 246 | |
| 49303 | 247 | 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 | 248 | |-> 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 | 249 | |> flat | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 250 | |> 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 | 251 | |> minimize_wits | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 252 | |> 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 | 253 | |
| 49714 | 254 |     fun wit_tac {context = ctxt, prems = _} =
 | 
| 51766 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 blanchet parents: 
51761diff
changeset | 255 | mk_comp_wit_tac ctxt (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 | 256 | (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 | 257 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 258 | val (bnf', lthy') = | 
| 51758 | 259 | bnf_def const_policy (K Dont_Note) qualify tacs wit_tac (SOME (oDs @ flat Dss)) Binding.empty | 
| 51767 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 blanchet parents: 
51766diff
changeset | 260 | Binding.empty [] (((((b, mapx), sets), bd), wits), SOME rel) lthy; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 261 | in | 
| 49503 | 262 | (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 | 263 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 264 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 265 | (* Killing live variables *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 266 | |
| 49502 | 267 | fun kill_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 268 | let | 
| 49425 | 269 | 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 | 270 | 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 | 271 | val dead = dead_of_bnf bnf; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 272 | 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 | 273 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 274 | (* TODO: check 0 < n <= live *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 275 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 276 | val (Ds, lthy1) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 277 | (Variable.invent_types (replicate dead HOLogic.typeS) lthy); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 278 | val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 279 | (Variable.invent_types (replicate live HOLogic.typeS) lthy1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 280 | val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 281 | (Variable.invent_types (replicate (live - n) HOLogic.typeS) lthy2); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 282 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 283 | val ((Asets, lives), _(*names_lthy*)) = lthy | 
| 49456 | 284 | |> 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 | 285 | ||>> mk_Frees "x" (drop n As); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 286 |     val xs = map (fn T => HOLogic.choice_const T $ absdummy T @{term True}) killedAs @ lives;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 287 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 288 | 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 | 289 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 290 | (*bnf.map id ... id*) | 
| 49303 | 291 | val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs); | 
| 49507 | 292 | (*bnf.rel (op =) ... (op =)*) | 
| 293 | val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs); | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 294 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 295 | val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf; | 
| 49303 | 296 | 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 | 297 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 298 | (*(|UNIV :: A1 set| +c ... +c |UNIV :: An set|) *c bnf.bd*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 299 | val bnf_bd = mk_bd_of_bnf Ds As bnf; | 
| 49303 | 300 | val bd = mk_cprod | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 301 | (Library.foldr1 (uncurry mk_csum) (map (mk_card_of o HOLogic.mk_UNIV) killedAs)) bnf_bd; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 302 | |
| 53270 | 303 | fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1; | 
| 53287 | 304 |     fun map_comp0_tac {context = ctxt, prems = _} =
 | 
| 53288 | 305 |       unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 306 | rtac refl 1; | 
| 51761 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 307 |     fun map_cong0_tac {context = ctxt, prems = _} =
 | 
| 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 308 | mk_kill_map_cong0_tac ctxt n (live - n) (map_cong0_of_bnf bnf); | 
| 53289 | 309 | val set_map0_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_map0_of_bnf bnf)); | 
| 49304 | 310 | fun bd_card_order_tac _ = mk_kill_bd_card_order_tac n (bd_card_order_of_bnf bnf); | 
| 311 | fun bd_cinfinite_tac _ = mk_kill_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf); | |
| 49303 | 312 | val set_bd_tacs = | 
| 49304 | 313 | map (fn thm => fn _ => mk_kill_set_bd_tac (bd_Card_order_of_bnf bnf) thm) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 314 | (drop n (set_bd_of_bnf bnf)); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 315 | |
| 49303 | 316 | val in_alt_thm = | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 317 | let | 
| 49303 | 318 | val inx = mk_in Asets sets T; | 
| 319 | val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T; | |
| 320 | 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 | 321 | in | 
| 51551 | 322 | Goal.prove_sorry lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 323 | end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 324 | |
| 49303 | 325 | fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 326 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 327 | fun rel_OO_Grp_tac _ = | 
| 49456 | 328 | let | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 329 | val rel_Grp = rel_Grp_of_bnf bnf RS sym | 
| 49463 | 330 | val thm = | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 331 |           (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 | 332 |             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 | 333 |               [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 | 334 | rel_conversep_of_bnf bnf RS sym], rel_Grp], | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 335 | trans OF [rel_OO_of_bnf bnf RS sym, rel_cong_of_bnf bnf OF | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 336 |                 (replicate n @{thm trans[OF Grp_UNIV_id[OF refl] eq_alt[symmetric]]} @
 | 
| 52660 | 337 |                  replicate (live - n) @{thm Grp_fst_snd})]]] RS sym);
 | 
| 49456 | 338 | in | 
| 49463 | 339 | rtac thm 1 | 
| 49456 | 340 | end; | 
| 341 | ||
| 53289 | 342 | val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52059diff
changeset | 343 | bd_cinfinite_tac set_bd_tacs map_wpull_tac rel_OO_Grp_tac; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 344 | |
| 49303 | 345 | 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 | 346 | |
| 49303 | 347 | val wits = map (fn t => fold absfree (Term.add_frees t []) t) | 
| 348 | (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 | 349 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 350 | fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 351 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 352 | val (bnf', lthy') = | 
| 51758 | 353 | bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME (killedAs @ Ds)) Binding.empty | 
| 51767 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 blanchet parents: 
51766diff
changeset | 354 | Binding.empty [] (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 355 | in | 
| 49503 | 356 | (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 | 357 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 358 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 359 | (* Adding dummy live variables *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 360 | |
| 49502 | 361 | fun lift_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 362 | let | 
| 49425 | 363 | 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 | 364 | 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 | 365 | val dead = dead_of_bnf bnf; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 366 | 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 | 367 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 368 | (* TODO: check 0 < n *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 369 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 370 | val (Ds, lthy1) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 371 | (Variable.invent_types (replicate dead HOLogic.typeS) lthy); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 372 | val ((newAs, As), lthy2) = apfst (chop n o map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 373 | (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 374 | val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 375 | (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy2); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 376 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 377 | val (Asets, _(*names_lthy*)) = lthy | 
| 49456 | 378 | |> 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 | 379 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 380 | 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 | 381 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 382 | (*%f1 ... fn. bnf.map*) | 
| 49303 | 383 | val mapx = | 
| 52923 | 384 | fold_rev Term.absdummy (map2 (curry op -->) newAs newBs) (mk_map_of_bnf Ds As Bs bnf); | 
| 49507 | 385 | (*%Q1 ... Qn. bnf.rel*) | 
| 386 | val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf); | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 387 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 388 | val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf; | 
| 49303 | 389 | 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 | 390 | |
| 49303 | 391 | 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 | 392 | |
| 53270 | 393 | fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1; | 
| 53287 | 394 |     fun map_comp0_tac {context = ctxt, prems = _} =
 | 
| 53288 | 395 |       unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 396 | rtac refl 1; | 
| 51761 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 397 |     fun map_cong0_tac {context = ctxt, prems = _} =
 | 
| 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 398 | rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); | 
| 53289 | 399 | val set_map0_tacs = | 
| 52059 | 400 | 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 | 401 | 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 | 402 | else | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 403 | replicate n (K empty_natural_tac) @ | 
| 53289 | 404 | map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf); | 
| 49303 | 405 | fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1; | 
| 406 | fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1; | |
| 407 | val set_bd_tacs = | |
| 52059 | 408 | 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 | 409 | 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 | 410 | else | 
| 49304 | 411 | replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @ | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 412 | (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf)); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 413 | |
| 49303 | 414 | val in_alt_thm = | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 415 | let | 
| 49303 | 416 | val inx = mk_in Asets sets T; | 
| 417 | val in_alt = mk_in (drop n Asets) bnf_sets T; | |
| 418 | 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 | 419 | in | 
| 51551 | 420 | Goal.prove_sorry lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 421 | end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 422 | |
| 49303 | 423 | fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 424 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 425 | fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm; | 
| 49456 | 426 | |
| 53289 | 427 | val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52059diff
changeset | 428 | bd_cinfinite_tac set_bd_tacs map_wpull_tac rel_OO_Grp_tac; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 429 | |
| 49303 | 430 | 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 | 431 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 432 | fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 433 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 434 | val (bnf', lthy') = | 
| 51767 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 blanchet parents: 
51766diff
changeset | 435 | bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty | 
| 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 blanchet parents: 
51766diff
changeset | 436 | [] (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 437 | in | 
| 49503 | 438 | (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 | 439 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 440 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 441 | (* 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 | 442 | |
| 49502 | 443 | fun permute_bnf qualify src dest bnf (unfold_set, lthy) = | 
| 444 | if src = dest then (bnf, (unfold_set, lthy)) else | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 445 | let | 
| 49425 | 446 | 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 | 447 | 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 | 448 | val dead = dead_of_bnf bnf; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 449 | val nwits = nwits_of_bnf bnf; | 
| 53040 | 450 | fun permute xs = permute_like (op =) src dest xs; | 
| 451 | fun unpermute xs = permute_like (op =) dest src xs; | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 452 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 453 | val (Ds, lthy1) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 454 | (Variable.invent_types (replicate dead HOLogic.typeS) lthy); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 455 | val (As, lthy2) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 456 | (Variable.invent_types (replicate live HOLogic.typeS) lthy1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 457 | val (Bs, _(*lthy3*)) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 458 | (Variable.invent_types (replicate live HOLogic.typeS) lthy2); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 459 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 460 | val (Asets, _(*names_lthy*)) = lthy | 
| 49456 | 461 | |> 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 | 462 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 463 | 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 | 464 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 465 | (*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*) | 
| 49303 | 466 | val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs)) | 
| 53038 | 467 | (Term.list_comb (mk_map_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0)))); | 
| 49507 | 468 | (*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*) | 
| 469 | val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs)) | |
| 53038 | 470 | (Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0)))); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 471 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 472 | val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf; | 
| 49303 | 473 | 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 | 474 | |
| 49303 | 475 | 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 | 476 | |
| 53270 | 477 | fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1; | 
| 53287 | 478 | fun map_comp0_tac _ = rtac (map_comp0_of_bnf bnf) 1; | 
| 51761 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 479 |     fun map_cong0_tac {context = ctxt, prems = _} =
 | 
| 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 blanchet parents: 
51758diff
changeset | 480 | rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); | 
| 53289 | 481 | val set_map0_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf)); | 
| 49303 | 482 | fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1; | 
| 483 | fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1; | |
| 484 | val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf)); | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 485 | |
| 49303 | 486 | val in_alt_thm = | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 487 | let | 
| 49303 | 488 | val inx = mk_in Asets sets T; | 
| 53038 | 489 | val in_alt = mk_in (unpermute Asets) bnf_sets T; | 
| 49303 | 490 | 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 | 491 | in | 
| 51551 | 492 | Goal.prove_sorry lthy [] [] goal (K (mk_permute_in_alt_tac src dest)) | 
| 49155 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 493 | |> Thm.close_derivation | 
| 
f51ab68f882f
do not trivialize important internal theorem in quick_and_dirty mode
 traytel parents: 
49123diff
changeset | 494 | end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 495 | |
| 49303 | 496 | fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 497 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 498 | fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm; | 
| 49456 | 499 | |
| 53289 | 500 | val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52059diff
changeset | 501 | bd_cinfinite_tac set_bd_tacs map_wpull_tac rel_OO_Grp_tac; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 502 | |
| 49303 | 503 | 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 | 504 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 505 | fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 506 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 507 | val (bnf', lthy') = | 
| 51767 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 blanchet parents: 
51766diff
changeset | 508 | bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty | 
| 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 blanchet parents: 
51766diff
changeset | 509 | [] (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 510 | in | 
| 49503 | 511 | (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 | 512 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 513 | |
| 49014 | 514 | (* Composition pipeline *) | 
| 515 | ||
| 516 | fun permute_and_kill qualify n src dest bnf = | |
| 517 | bnf | |
| 518 | |> permute_bnf qualify src dest | |
| 49304 | 519 | #> uncurry (kill_bnf qualify n); | 
| 49014 | 520 | |
| 521 | fun lift_and_permute qualify n src dest bnf = | |
| 522 | bnf | |
| 49304 | 523 | |> lift_bnf qualify n | 
| 49014 | 524 | #> uncurry (permute_bnf qualify src dest); | 
| 525 | ||
| 49502 | 526 | fun normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy = | 
| 49014 | 527 | let | 
| 528 | val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass; | |
| 52985 | 529 | val kill_poss = map (find_indices op = Ds) Ass; | 
| 530 | val live_poss = map2 (subtract op =) kill_poss before_kill_src; | |
| 49014 | 531 | val before_kill_dest = map2 append kill_poss live_poss; | 
| 532 | val kill_ns = map length kill_poss; | |
| 49502 | 533 | val (inners', (unfold_set', lthy')) = | 
| 49014 | 534 | fold_map5 (fn i => permute_and_kill (qualify i)) | 
| 535 | (if length bnfs = 1 then [0] else (1 upto length bnfs)) | |
| 49502 | 536 | kill_ns before_kill_src before_kill_dest bnfs (unfold_set, lthy); | 
| 49014 | 537 | |
| 538 | val Ass' = map2 (map o nth) Ass live_poss; | |
| 539 | val As = sort Ass'; | |
| 540 | val after_lift_dest = replicate (length Ass') (0 upto (length As - 1)); | |
| 541 | val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass'; | |
| 52985 | 542 | val new_poss = map2 (subtract op =) old_poss after_lift_dest; | 
| 49014 | 543 | val after_lift_src = map2 append new_poss old_poss; | 
| 544 | val lift_ns = map (fn xs => length As - length xs) Ass'; | |
| 545 | in | |
| 546 | ((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i)) | |
| 547 | (if length bnfs = 1 then [0] else (1 upto length bnfs)) | |
| 49502 | 548 | lift_ns after_lift_src after_lift_dest inners' (unfold_set', lthy')) | 
| 49014 | 549 | end; | 
| 550 | ||
| 551 | fun default_comp_sort Ass = | |
| 552 | Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []); | |
| 553 | ||
| 49502 | 554 | fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (unfold_set, lthy) = | 
| 49014 | 555 | let | 
| 49425 | 556 | val b = name_of_bnf outer; | 
| 49014 | 557 | |
| 49121 | 558 | val Ass = map (map Term.dest_TFree) tfreess; | 
| 49014 | 559 | val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) []; | 
| 560 | ||
| 49502 | 561 | val ((kill_poss, As), (inners', (unfold_set', lthy'))) = | 
| 562 | normalize_bnfs qualify Ass Ds sort inners unfold_set lthy; | |
| 49014 | 563 | |
| 564 | val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss); | |
| 565 | val As = map TFree As; | |
| 566 | in | |
| 49425 | 567 | apfst (rpair (Ds, As)) | 
| 49502 | 568 | (clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy')) | 
| 49014 | 569 | end; | 
| 570 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 571 | (* 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 | 572 | |
| 53264 | 573 | fun seal_bnf qualify (unfold_set : unfold_set) b Ds bnf lthy = | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 574 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 575 | 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 | 576 | 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 | 577 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 578 | val (As, lthy1) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 579 | (Variable.invent_types (replicate live HOLogic.typeS) (fold Variable.declare_typ Ds lthy)); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 580 | val (Bs, _) = apfst (map TFree) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 581 | (Variable.invent_types (replicate live HOLogic.typeS) lthy1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 582 | |
| 49713 
3d07ddf70f8b
do not expose details of internal data structures for composition of BNFs
 traytel parents: 
49669diff
changeset | 583 | val map_unfolds = #map_unfolds unfold_set; | 
| 
3d07ddf70f8b
do not expose details of internal data structures for composition of BNFs
 traytel parents: 
49669diff
changeset | 584 | val set_unfoldss = #set_unfoldss unfold_set; | 
| 
3d07ddf70f8b
do not expose details of internal data structures for composition of BNFs
 traytel parents: 
49669diff
changeset | 585 | val rel_unfolds = #rel_unfolds unfold_set; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 586 | |
| 49507 | 587 | val expand_maps = | 
| 588 | fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds); | |
| 589 | val expand_sets = | |
| 590 | fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss); | |
| 591 | val expand_rels = | |
| 592 | fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) rel_unfolds); | |
| 49504 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 blanchet parents: 
49503diff
changeset | 593 | val unfold_maps = fold (unfold_thms lthy o single) map_unfolds; | 
| 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 blanchet parents: 
49503diff
changeset | 594 | val unfold_sets = fold (unfold_thms lthy) set_unfoldss; | 
| 49507 | 595 | val unfold_rels = unfold_thms lthy rel_unfolds; | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51837diff
changeset | 596 | val unfold_all = unfold_sets o unfold_maps o unfold_rels; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 597 | val bnf_map = expand_maps (mk_map_of_bnf Ds As Bs bnf); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 598 | val bnf_sets = map (expand_maps o expand_sets) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 599 | (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 | 600 | val bnf_bd = mk_bd_of_bnf Ds As bnf; | 
| 49507 | 601 | val bnf_rel = expand_rels (mk_rel_of_bnf Ds As Bs bnf); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 602 | 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 | 603 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 604 | (*bd should only depend on dead type variables!*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 605 | val bd_repT = fst (dest_relT (fastype_of bnf_bd)); | 
| 53264 | 606 |     val bdT_bind = qualify (Binding.suffix_name ("_" ^ bdTN) b);
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 607 | val params = fold Term.add_tfreesT Ds []; | 
| 49185 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 traytel parents: 
49155diff
changeset | 608 | val deads = map TFree params; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 609 | |
| 49228 | 610 | val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) = | 
| 49835 | 611 | typedef (bdT_bind, params, NoSyn) | 
| 49228 | 612 | (HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 613 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 614 | val bnf_bd' = mk_dir_image bnf_bd | 
| 49185 
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
 traytel parents: 
49155diff
changeset | 615 | (Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, deads))) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 616 | |
| 49228 | 617 | val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info); | 
| 618 | val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj (#Abs_cases bdT_loc_info); | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 619 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 620 |     val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf];
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 621 | val bd_card_order = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 622 |       @{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf];
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 623 | val bd_cinfinite = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 624 |       (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 625 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 626 | val set_bds = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 627 |       map (fn thm => @{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso]) (set_bd_of_bnf bnf);
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 628 | |
| 49463 | 629 |     fun mk_tac thm {context = ctxt, prems = _} =
 | 
| 630 | (rtac (unfold_all thm) THEN' | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 631 | SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1; | 
| 49456 | 632 | |
| 53287 | 633 | val tacs = zip_axioms (mk_tac (map_id0_of_bnf bnf)) (mk_tac (map_comp0_of_bnf bnf)) | 
| 53289 | 634 | (mk_tac (map_cong0_of_bnf bnf)) (map mk_tac (set_map0_of_bnf bnf)) | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52059diff
changeset | 635 | (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds) | 
| 49463 | 636 | (mk_tac (map_wpull_of_bnf bnf)) | 
| 51895 
e0bf21531ed5
do not unfold the definition of the relator as it is not defined in terms of srel anymore
 traytel parents: 
51893diff
changeset | 637 | (mk_tac (rel_OO_Grp_of_bnf bnf)); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 638 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 639 | val bnf_wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 640 | |
| 49463 | 641 | fun wit_tac _ = mk_simple_wit_tac (map unfold_all (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 | 642 | |
| 53264 | 643 | val (bnf', lthy') = | 
| 644 | bnf_def Hardly_Inline (user_policy Dont_Note) qualify tacs wit_tac (SOME deads) | |
| 645 | Binding.empty Binding.empty [] | |
| 646 | (((((b, bnf_map), bnf_sets), Term.absdummy T bnf_bd'), bnf_wits), SOME bnf_rel) lthy; | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 647 | in | 
| 49456 | 648 | ((bnf', deads), lthy') | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 649 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 650 | |
| 53222 | 651 | exception BAD_DEAD of typ * typ; | 
| 652 | ||
| 653 | fun bnf_of_typ _ _ _ _ (T as TFree _) accum = ((ID_bnf, ([], [T])), accum) | |
| 654 | | bnf_of_typ _ _ _ _ (TVar _) _ = error "Unexpected schematic variable" | |
| 655 | | bnf_of_typ const_policy qualify' sort Xs (T as Type (C, Ts)) (unfold_set, lthy) = | |
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 656 | let | 
| 53222 | 657 | fun check_bad_dead ((_, (deads, _)), _) = | 
| 658 | let val Ds = fold Term.add_tfreesT deads [] in | |
| 659 | (case Library.inter (op =) Ds Xs of [] => () | |
| 660 | | X :: _ => raise BAD_DEAD (TFree X, T)) | |
| 661 | end; | |
| 662 | ||
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 663 | val tfrees = Term.add_tfreesT T []; | 
| 49236 | 664 | 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 | 665 | in | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 666 | (case bnf_opt of | 
| 49502 | 667 | NONE => ((DEADID_bnf, ([T], [])), (unfold_set, lthy)) | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 668 | | SOME bnf => | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 669 | if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 670 | let | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 671 | val T' = T_of_bnf bnf; | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 672 | val deads = deads_of_bnf bnf; | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 673 | val lives = lives_of_bnf bnf; | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 674 | val tvars' = Term.add_tvarsT T' []; | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 675 | val deads_lives = | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 676 | pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees))) | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 677 | (deads, lives); | 
| 49502 | 678 | in ((bnf, deads_lives), (unfold_set, lthy)) end | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 679 | else | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 680 | let | 
| 49425 | 681 | val name = Long_Name.base_name C; | 
| 682 | fun qualify i = | |
| 683 | let val namei = name ^ nonzero_string_of_int i; | |
| 684 | 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 | 685 | val odead = dead_of_bnf bnf; | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 686 | val olive = live_of_bnf bnf; | 
| 52985 | 687 |             val oDs_pos = find_indices op = [TFree ("dead", [])] (snd (Term.dest_Type
 | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 688 |               (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
 | 
| 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 689 | 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 | 690 | val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1)); | 
| 49502 | 691 | val ((inners, (Dss, Ass)), (unfold_set', lthy')) = | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 692 | apfst (apsnd split_list o split_list) | 
| 53222 | 693 | (fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort Xs) | 
| 49502 | 694 | (if length Ts' = 1 then [0] else (1 upto length Ts')) Ts' (unfold_set, lthy)); | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 695 | in | 
| 49502 | 696 | compose_bnf const_policy qualify sort bnf inners oDs Dss Ass (unfold_set', lthy') | 
| 49186 
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
 traytel parents: 
49185diff
changeset | 697 | end) | 
| 53222 | 698 | |> tap check_bad_dead | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 699 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 700 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 701 | end; |