src/HOL/BNF/Tools/bnf_comp.ML
author wenzelm
Fri, 12 Oct 2012 15:08:29 +0200
changeset 49833 1d80798e8d8a
parent 49714 2c7c32f34220
child 49835 31f32ec4d766
permissions -rw-r--r--
discontinued typedef with implicit set_def;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49507
diff 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
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    11
  val ID_bnf: BNF_Def.BNF
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    12
  val DEADID_bnf: BNF_Def.BNF
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    13
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    14
  type unfold_set
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    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
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
    17
  val bnf_of_typ: BNF_Def.const_policy -> (binding -> binding) ->
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    18
    ((string * sort) list list -> (string * sort) list) -> typ -> unfold_set * Proof.context ->
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    19
    (BNF_Def.BNF * (typ list * typ list)) * (unfold_set * Proof.context)
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
    20
  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
    21
  val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    22
    (''a list list -> ''a list) -> BNF_Def.BNF list -> unfold_set -> Proof.context ->
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    23
    (int list list * ''a list) * (BNF_Def.BNF list * (unfold_set * Proof.context))
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    24
  val seal_bnf: unfold_set -> binding -> typ list -> BNF_Def.BNF -> Proof.context ->
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49155
diff changeset
    25
    (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
    26
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
structure BNF_Comp : BNF_COMP =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
struct
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
open BNF_Def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
open BNF_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
open BNF_Tactics
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
open BNF_Comp_Tactics
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    36
val ID_bnf = the (bnf_of @{context} "Basic_BNFs.ID");
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    37
val DEADID_bnf = the (bnf_of @{context} "Basic_BNFs.DEADID");
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    38
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49538
diff changeset
    39
(* TODO: Replace by "BNF_Defs.defs list" *)
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    40
type unfold_set = {
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  map_unfolds: thm list,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
  set_unfoldss: thm list list,
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    43
  rel_unfolds: thm list,
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
    44
  srel_unfolds: thm list
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    47
val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = [], srel_unfolds = []};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
    49
fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new;
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
    50
fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms;
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
    51
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    52
fun add_to_unfolds map sets rel srel
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    53
  {map_unfolds, set_unfoldss, rel_unfolds, srel_unfolds} =
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
    54
  {map_unfolds = add_to_thms map_unfolds map,
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
    55
    set_unfoldss = adds_to_thms set_unfoldss sets,
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    56
    rel_unfolds = add_to_thms rel_unfolds rel,
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
    57
    srel_unfolds = add_to_thms srel_unfolds srel};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
    59
fun add_bnf_to_unfolds bnf =
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    60
  add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf)
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
    61
    (srel_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
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
    65
fun mk_killN n = "_kill" ^ string_of_int n;
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
    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
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
    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
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
    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
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   103
    val ((((fs', Qs'), Asets), xs), _(*names_lthy*)) = lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
      |> apfst snd o mk_Frees' "f" (map2 (curry (op -->)) As Bs)
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   105
      ||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs)
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   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
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   122
    (*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*)
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   123
    val rel = fold_rev Term.abs Qs'
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   124
      (Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer,
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   125
        map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
49630
9f6ca87ab405 tuned tactics
traytel
parents: 49586
diff changeset
   152
    fun map_id_tac _ =
9f6ca87ab405 tuned tactics
traytel
parents: 49586
diff changeset
   153
      mk_comp_map_id_tac (map_id_of_bnf outer) (map_cong_of_bnf outer)
9f6ca87ab405 tuned tactics
traytel
parents: 49586
diff changeset
   154
        (map map_id_of_bnf inners);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
49303
blanchet
parents: 49236
diff changeset
   156
    fun map_comp_tac _ =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
      mk_comp_map_comp_tac (map_comp_of_bnf outer) (map_cong_of_bnf outer)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
        (map map_comp_of_bnf inners);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
49303
blanchet
parents: 49236
diff changeset
   160
    fun mk_single_set_natural_tac i _ =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   161
      mk_comp_set_natural_tac (map_comp_of_bnf outer) (map_cong_of_bnf outer)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
        (collect_set_natural_of_bnf outer)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
        (map ((fn thms => nth thms i) o set_natural_of_bnf) inners);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
49303
blanchet
parents: 49236
diff changeset
   165
    val set_natural_tacs = map mk_single_set_natural_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
blanchet
parents: 49236
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
blanchet
parents: 49236
diff changeset
   173
    val set_alt_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
      if ! quick_and_dirty then
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   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: 49019
diff changeset
   177
        map (fn goal =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49019
diff changeset
   178
          Skip_Proof.prove lthy [] [] goal
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   179
            (fn {context = ctxt, prems = _} =>
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   180
              mk_comp_set_alt_tac ctxt (collect_set_natural_of_bnf outer))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49019
diff changeset
   181
          |> Thm.close_derivation)
49303
blanchet
parents: 49236
diff changeset
   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
49303
blanchet
parents: 49236
diff changeset
   184
    fun map_cong_tac _ =
blanchet
parents: 49236
diff changeset
   185
      mk_comp_map_cong_tac set_alt_thms (map_cong_of_bnf outer) (map map_cong_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
blanchet
parents: 49236
diff changeset
   187
    val set_bd_tacs =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
      if ! quick_and_dirty then
49669
620fa6272c48 fixed quick-and-dirty mode
blanchet
parents: 49630
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
blanchet
parents: 49236
diff changeset
   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
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   201
          map2 (fn set_alt => fn single_set_bds => fn {context = ctxt, prems = _} =>
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   202
            mk_comp_set_bd_tac ctxt set_alt single_set_bds)
49303
blanchet
parents: 49236
diff changeset
   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
blanchet
parents: 49236
diff changeset
   206
    val in_alt_thm =
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   207
      let
49303
blanchet
parents: 49236
diff changeset
   208
        val inx = mk_in Asets sets CCA;
blanchet
parents: 49236
diff changeset
   209
        val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;
blanchet
parents: 49236
diff changeset
   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: 49123
diff changeset
   211
      in
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   212
        Skip_Proof.prove lthy [] [] goal
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   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: 49123
diff changeset
   214
        |> Thm.close_derivation
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff 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
blanchet
parents: 49236
diff changeset
   217
    fun in_bd_tac _ =
blanchet
parents: 49236
diff changeset
   218
      mk_comp_in_bd_tac in_alt_thm (map in_bd_of_bnf inners) (in_bd_of_bnf outer)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
        (map bd_Cinfinite_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
   220
49303
blanchet
parents: 49236
diff changeset
   221
    fun map_wpull_tac _ =
blanchet
parents: 49236
diff changeset
   222
      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
   223
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   224
    fun srel_O_Gr_tac _ =
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   225
      let
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   226
        val basic_thms = @{thms mem_Collect_eq fst_conv snd_conv}; (*TODO: tune*)
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   227
        val outer_srel_Gr = srel_Gr_of_bnf outer RS sym;
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   228
        val outer_srel_cong = srel_cong_of_bnf outer;
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   229
        val thm =
49512
82d99fe04018 clean up lemmas used for composition
blanchet
parents: 49510
diff changeset
   230
          (trans OF [in_alt_thm RS @{thm O_Gr_cong},
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   231
             trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   232
               [trans OF [outer_srel_Gr RS @{thm arg_cong[of _ _ converse]},
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   233
                 srel_converse_of_bnf outer RS sym], outer_srel_Gr],
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   234
               trans OF [srel_O_of_bnf outer RS sym, outer_srel_cong OF
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   235
                 (map (fn bnf => srel_O_Gr_of_bnf bnf RS sym) inners)]]] RS sym)
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   236
          |> unfold_thms lthy (basic_thms @ srel_def_of_bnf outer :: map srel_def_of_bnf inners);
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   237
      in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49503
diff changeset
   238
        unfold_thms_tac lthy basic_thms THEN rtac thm 1
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   239
      end;
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   240
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   241
    val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   242
      bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
    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
   245
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
    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
   247
      (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
   248
        Dss inwitss inners);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
    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
   251
49303
blanchet
parents: 49236
diff changeset
   252
    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
   253
      |-> 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
   254
      |> flat
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
      |> 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
   256
      |> minimize_wits
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
      |> 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
   258
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   259
    fun wit_tac {context = ctxt, prems = _} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   260
      mk_comp_wit_tac ctxt (wit_thms_of_bnf outer) (collect_set_natural_of_bnf outer)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
        (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
   262
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
    val (bnf', lthy') =
49538
c90818b63599 simplified fact policies
blanchet
parents: 49512
diff changeset
   264
      bnf_def const_policy (K Dont_Note) qualify tacs wit_tac (SOME (oDs @ flat Dss))
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   265
        (((((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
   266
  in
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
   267
    (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
   268
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
(* Killing live variables *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   272
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
   273
  let
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   274
    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
   275
    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
   276
    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
   277
    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
   278
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
    (* TODO: check 0 < n <= live *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   280
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
    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
   282
      (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
   283
    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
   284
      (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
   285
    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
   286
      (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
   287
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
    val ((Asets, lives), _(*names_lthy*)) = lthy
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   289
      |> 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
   290
      ||>> 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
   291
    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
   292
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
    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
   294
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
    (*bnf.map id ... id*)
49303
blanchet
parents: 49236
diff changeset
   296
    val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   297
    (*bnf.rel (op =) ... (op =)*)
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   298
    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
   299
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
49303
blanchet
parents: 49236
diff changeset
   301
    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
   302
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
    (*(|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
   304
    val bnf_bd = mk_bd_of_bnf Ds As bnf;
49303
blanchet
parents: 49236
diff changeset
   305
    val bd = mk_cprod
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
      (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
   307
49303
blanchet
parents: 49236
diff changeset
   308
    fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   309
    fun map_comp_tac {context = ctxt, prems = _} =
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   310
      unfold_thms_tac ctxt ((map_comp_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
   311
      rtac refl 1;
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   312
    fun map_cong_tac {context = ctxt, prems = _} =
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   313
      mk_kill_map_cong_tac ctxt n (live - n) (map_cong_of_bnf bnf);
49303
blanchet
parents: 49236
diff changeset
   314
    val set_natural_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_natural_of_bnf bnf));
49304
blanchet
parents: 49303
diff changeset
   315
    fun bd_card_order_tac _ = mk_kill_bd_card_order_tac n (bd_card_order_of_bnf bnf);
blanchet
parents: 49303
diff changeset
   316
    fun bd_cinfinite_tac _ = mk_kill_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf);
49303
blanchet
parents: 49236
diff changeset
   317
    val set_bd_tacs =
49304
blanchet
parents: 49303
diff changeset
   318
      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
   319
        (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
   320
49303
blanchet
parents: 49236
diff changeset
   321
    val in_alt_thm =
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   322
      let
49303
blanchet
parents: 49236
diff changeset
   323
        val inx = mk_in Asets sets T;
blanchet
parents: 49236
diff changeset
   324
        val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;
blanchet
parents: 49236
diff changeset
   325
        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   326
      in
49304
blanchet
parents: 49303
diff changeset
   327
        Skip_Proof.prove lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   328
      end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   329
49303
blanchet
parents: 49236
diff changeset
   330
    fun in_bd_tac _ =
49304
blanchet
parents: 49303
diff changeset
   331
      mk_kill_in_bd_tac n (live > n) in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf)
49303
blanchet
parents: 49236
diff changeset
   332
        (bd_Cinfinite_of_bnf bnf) (bd_Cnotzero_of_bnf bnf);
blanchet
parents: 49236
diff changeset
   333
    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
   334
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   335
    fun srel_O_Gr_tac _ =
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   336
      let
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   337
        val srel_Gr = srel_Gr_of_bnf bnf RS sym
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   338
        val thm =
49512
82d99fe04018 clean up lemmas used for composition
blanchet
parents: 49510
diff changeset
   339
          (trans OF [in_alt_thm RS @{thm O_Gr_cong},
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   340
            trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   341
              [trans OF [srel_Gr RS @{thm arg_cong[of _ _ converse]},
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   342
                srel_converse_of_bnf bnf RS sym], srel_Gr],
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   343
              trans OF [srel_O_of_bnf bnf RS sym, srel_cong_of_bnf bnf OF
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   344
                (replicate n @{thm trans[OF Gr_UNIV_id[OF refl] Id_alt[symmetric]]} @
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   345
                 replicate (live - n) @{thm Gr_fst_snd})]]] RS sym)
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   346
          |> unfold_thms lthy (srel_def_of_bnf bnf :: @{thms Id_def' mem_Collect_eq split_conv});
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   347
      in
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   348
        rtac thm 1
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   349
      end;
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   350
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   351
    val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   352
      bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   353
49303
blanchet
parents: 49236
diff changeset
   354
    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
   355
49303
blanchet
parents: 49236
diff changeset
   356
    val wits = map (fn t => fold absfree (Term.add_frees t []) t)
blanchet
parents: 49236
diff changeset
   357
      (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
   358
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   359
    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
   360
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
    val (bnf', lthy') =
49538
c90818b63599 simplified fact policies
blanchet
parents: 49512
diff changeset
   362
      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME (killedAs @ Ds))
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   363
        (((((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
   364
  in
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
   365
    (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
   366
  end;
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
(* Adding dummy live variables *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   369
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   370
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
   371
  let
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   372
    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
   373
    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
   374
    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
   375
    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
   376
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   377
    (* TODO: check 0 < n *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   378
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
    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
   380
      (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
   381
    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
   382
      (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
   383
    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
   384
      (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
   385
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   386
    val (Asets, _(*names_lthy*)) = lthy
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   387
      |> 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
   388
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   389
    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
   390
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   391
    (*%f1 ... fn. bnf.map*)
49303
blanchet
parents: 49236
diff changeset
   392
    val mapx =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   393
      fold_rev Term.absdummy (map2 (curry (op -->)) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   394
    (*%Q1 ... Qn. bnf.rel*)
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   395
    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
   396
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   397
    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
49303
blanchet
parents: 49236
diff changeset
   398
    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
   399
49303
blanchet
parents: 49236
diff changeset
   400
    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
   401
49303
blanchet
parents: 49236
diff changeset
   402
    fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   403
    fun map_comp_tac {context = ctxt, prems = _} =
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   404
      unfold_thms_tac ctxt ((map_comp_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
   405
      rtac refl 1;
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   406
    fun map_cong_tac {context = ctxt, prems = _} =
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   407
      rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);
49303
blanchet
parents: 49236
diff changeset
   408
    val set_natural_tacs =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
      if ! quick_and_dirty then
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
        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
   411
      else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   412
        replicate n (K empty_natural_tac) @
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   413
        map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf);
49303
blanchet
parents: 49236
diff changeset
   414
    fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
blanchet
parents: 49236
diff changeset
   415
    fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
blanchet
parents: 49236
diff changeset
   416
    val set_bd_tacs =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   417
      if ! quick_and_dirty then
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
        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
   419
      else
49304
blanchet
parents: 49303
diff changeset
   420
        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
   421
        (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
   422
49303
blanchet
parents: 49236
diff changeset
   423
    val in_alt_thm =
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   424
      let
49303
blanchet
parents: 49236
diff changeset
   425
        val inx = mk_in Asets sets T;
blanchet
parents: 49236
diff changeset
   426
        val in_alt = mk_in (drop n Asets) bnf_sets T;
blanchet
parents: 49236
diff changeset
   427
        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   428
      in
49304
blanchet
parents: 49303
diff changeset
   429
        Skip_Proof.prove lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   430
      end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   431
49304
blanchet
parents: 49303
diff changeset
   432
    fun in_bd_tac _ = mk_lift_in_bd_tac n in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
49303
blanchet
parents: 49236
diff changeset
   433
    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
   434
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   435
    fun srel_O_Gr_tac _ =
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   436
      mk_simple_srel_O_Gr_tac lthy (srel_def_of_bnf bnf) (srel_O_Gr_of_bnf bnf) in_alt_thm;
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   437
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   438
    val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   439
      bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   440
49303
blanchet
parents: 49236
diff changeset
   441
    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
   442
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   443
    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
   444
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   445
    val (bnf', lthy') =
49538
c90818b63599 simplified fact policies
blanchet
parents: 49512
diff changeset
   446
      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   447
        (((((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
   448
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   449
  in
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
   450
    (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
   451
  end;
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
(* 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
   454
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   455
fun permute_bnf qualify src dest bnf (unfold_set, lthy) =
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   456
  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
   457
  let
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   458
    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
   459
    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
   460
    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
   461
    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
   462
    fun permute xs = mk_permute src dest xs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   463
    fun permute_rev xs = mk_permute dest src xs;
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
    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
   466
      (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
   467
    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
   468
      (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
   469
    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
   470
      (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
   471
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   472
    val (Asets, _(*names_lthy*)) = lthy
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   473
      |> 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
   474
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   475
    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
   476
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   477
    (*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*)
49303
blanchet
parents: 49236
diff changeset
   478
    val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs))
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   479
      (Term.list_comb (mk_map_of_bnf Ds As Bs bnf, permute_rev (map Bound (live - 1 downto 0))));
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   480
    (*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*)
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   481
    val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs))
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   482
      (Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, permute_rev (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
   483
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   484
    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
49303
blanchet
parents: 49236
diff changeset
   485
    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
   486
49303
blanchet
parents: 49236
diff changeset
   487
    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
   488
49303
blanchet
parents: 49236
diff changeset
   489
    fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
blanchet
parents: 49236
diff changeset
   490
    fun map_comp_tac _ = rtac (map_comp_of_bnf bnf) 1;
49714
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   491
    fun map_cong_tac {context = ctxt, prems = _} =
2c7c32f34220 made SML/NJ happier
traytel
parents: 49713
diff changeset
   492
      rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);
49303
blanchet
parents: 49236
diff changeset
   493
    val set_natural_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf));
blanchet
parents: 49236
diff changeset
   494
    fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
blanchet
parents: 49236
diff changeset
   495
    fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
blanchet
parents: 49236
diff changeset
   496
    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
   497
49303
blanchet
parents: 49236
diff changeset
   498
    val in_alt_thm =
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   499
      let
49303
blanchet
parents: 49236
diff changeset
   500
        val inx = mk_in Asets sets T;
blanchet
parents: 49236
diff changeset
   501
        val in_alt = mk_in (permute_rev Asets) bnf_sets T;
blanchet
parents: 49236
diff changeset
   502
        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
49155
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   503
      in
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   504
        Skip_Proof.prove lthy [] [] goal (K (mk_permute_in_alt_tac src dest))
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   505
        |> Thm.close_derivation
f51ab68f882f do not trivialize important internal theorem in quick_and_dirty mode
traytel
parents: 49123
diff changeset
   506
      end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   507
49303
blanchet
parents: 49236
diff changeset
   508
    fun in_bd_tac _ =
blanchet
parents: 49236
diff changeset
   509
      mk_permute_in_bd_tac src dest in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
blanchet
parents: 49236
diff changeset
   510
    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
   511
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   512
    fun srel_O_Gr_tac _ =
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   513
      mk_simple_srel_O_Gr_tac lthy (srel_def_of_bnf bnf) (srel_O_Gr_of_bnf bnf) in_alt_thm;
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   514
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   515
    val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   516
      bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   517
49303
blanchet
parents: 49236
diff changeset
   518
    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
   519
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   520
    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
   521
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   522
    val (bnf', lthy') =
49538
c90818b63599 simplified fact policies
blanchet
parents: 49512
diff changeset
   523
      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   524
        (((((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
   525
  in
49503
dbbde075a42e simplified code
blanchet
parents: 49502
diff changeset
   526
    (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
   527
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   528
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   529
(* Composition pipeline *)
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   530
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   531
fun permute_and_kill qualify n src dest bnf =
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   532
  bnf
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   533
  |> permute_bnf qualify src dest
49304
blanchet
parents: 49303
diff changeset
   534
  #> uncurry (kill_bnf qualify n);
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   535
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   536
fun lift_and_permute qualify n src dest bnf =
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   537
  bnf
49304
blanchet
parents: 49303
diff changeset
   538
  |> lift_bnf qualify n
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   539
  #> uncurry (permute_bnf qualify src dest);
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   540
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   541
fun normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy =
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   542
  let
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   543
    val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   544
    val kill_poss = map (find_indices Ds) Ass;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   545
    val live_poss = map2 (subtract (op =)) kill_poss before_kill_src;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   546
    val before_kill_dest = map2 append kill_poss live_poss;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   547
    val kill_ns = map length kill_poss;
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   548
    val (inners', (unfold_set', lthy')) =
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   549
      fold_map5 (fn i => permute_and_kill (qualify i))
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   550
        (if length bnfs = 1 then [0] else (1 upto length bnfs))
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   551
        kill_ns before_kill_src before_kill_dest bnfs (unfold_set, lthy);
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   552
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   553
    val Ass' = map2 (map o nth) Ass live_poss;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   554
    val As = sort Ass';
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   555
    val after_lift_dest = replicate (length Ass') (0 upto (length As - 1));
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   556
    val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass';
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   557
    val new_poss = map2 (subtract (op =)) old_poss after_lift_dest;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   558
    val after_lift_src = map2 append new_poss old_poss;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   559
    val lift_ns = map (fn xs => length As - length xs) Ass';
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   560
  in
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   561
    ((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i))
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   562
      (if length bnfs = 1 then [0] else (1 upto length bnfs))
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   563
      lift_ns after_lift_src after_lift_dest inners' (unfold_set', lthy'))
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   564
  end;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   565
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   566
fun default_comp_sort Ass =
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   567
  Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []);
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   568
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   569
fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (unfold_set, lthy) =
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   570
  let
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   571
    val b = name_of_bnf outer;
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   572
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49109
diff changeset
   573
    val Ass = map (map Term.dest_TFree) tfreess;
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   574
    val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   575
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   576
    val ((kill_poss, As), (inners', (unfold_set', lthy'))) =
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   577
      normalize_bnfs qualify Ass Ds sort inners unfold_set lthy;
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   578
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   579
    val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss);
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   580
    val As = map TFree As;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   581
  in
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   582
    apfst (rpair (Ds, As))
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   583
      (clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy'))
49014
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   584
  end;
332e5e661675 have "bnf_of_typ" command seal the BNF
blanchet
parents: 48975
diff changeset
   585
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   586
(* 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
   587
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   588
fun seal_bnf 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
   589
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   590
    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
   591
    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
   592
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   593
    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
   594
      (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
   595
    val (Bs, _) = apfst (map TFree)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   596
      (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
   597
49713
3d07ddf70f8b do not expose details of internal data structures for composition of BNFs
traytel
parents: 49669
diff changeset
   598
    val map_unfolds = #map_unfolds unfold_set;
3d07ddf70f8b do not expose details of internal data structures for composition of BNFs
traytel
parents: 49669
diff changeset
   599
    val set_unfoldss = #set_unfoldss unfold_set;
3d07ddf70f8b do not expose details of internal data structures for composition of BNFs
traytel
parents: 49669
diff changeset
   600
    val rel_unfolds = #rel_unfolds unfold_set;
3d07ddf70f8b do not expose details of internal data structures for composition of BNFs
traytel
parents: 49669
diff changeset
   601
    val srel_unfolds = #srel_unfolds unfold_set;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   602
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   603
    val expand_maps =
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   604
      fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds);
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   605
    val expand_sets =
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   606
      fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss);
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   607
    val expand_rels =
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   608
      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: 49503
diff changeset
   609
    val unfold_maps = fold (unfold_thms lthy o single) map_unfolds;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49503
diff changeset
   610
    val unfold_sets = fold (unfold_thms lthy) set_unfoldss;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   611
    val unfold_rels = unfold_thms lthy rel_unfolds;
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   612
    val unfold_srels = unfold_thms lthy srel_unfolds;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   613
    val unfold_all = unfold_sets o unfold_maps o unfold_rels o unfold_srels;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   614
    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
   615
    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
   616
      (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
   617
    val bnf_bd = mk_bd_of_bnf Ds As bnf;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   618
    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
   619
    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
   620
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   621
    (*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
   622
    val bd_repT = fst (dest_relT (fastype_of bnf_bd));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   623
    val bdT_bind = Binding.suffix_name ("_" ^ bdTN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   624
    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: 49155
diff changeset
   625
    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
   626
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49186
diff changeset
   627
    val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) =
49833
1d80798e8d8a discontinued typedef with implicit set_def;
wenzelm
parents: 49714
diff changeset
   628
      typedef NONE (bdT_bind, params, NoSyn)
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49186
diff changeset
   629
        (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
   630
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   631
    val bnf_bd' = mk_dir_image bnf_bd
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49155
diff changeset
   632
      (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
   633
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49186
diff changeset
   634
    val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info);
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49186
diff changeset
   635
    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
   636
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   637
    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
   638
    val bd_card_order =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   639
      @{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
   640
    val bd_cinfinite =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   641
      (@{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
   642
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   643
    val set_bds =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   644
      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
   645
    val in_bd =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   646
      @{thm ordLeq_ordIso_trans} OF [in_bd_of_bnf bnf,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   647
        @{thm cexp_cong2_Cnotzero} OF [bd_ordIso, if live = 0 then
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   648
          @{thm ctwo_Cnotzero} else @{thm ctwo_Cnotzero} RS @{thm csum_Cnotzero2},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   649
            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
   650
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   651
    fun mk_tac thm {context = ctxt, prems = _} =
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   652
      (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
   653
      SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   654
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   655
    val tacs = zip_axioms (mk_tac (map_id_of_bnf bnf)) (mk_tac (map_comp_of_bnf bnf))
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   656
      (mk_tac (map_cong_of_bnf bnf)) (map mk_tac (set_natural_of_bnf bnf))
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   657
      (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds) (mk_tac in_bd)
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   658
      (mk_tac (map_wpull_of_bnf bnf))
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   659
      (mk_tac (unfold_thms lthy [srel_def_of_bnf bnf] (srel_O_Gr_of_bnf bnf)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   660
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   661
    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
   662
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   663
    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
   664
49538
c90818b63599 simplified fact policies
blanchet
parents: 49512
diff changeset
   665
    val (bnf', lthy') = bnf_def Hardly_Inline (user_policy Dont_Note) I tacs wit_tac (SOME deads)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   666
      (((((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
   667
  in
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   668
    ((bnf', deads), lthy')
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   669
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   670
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   671
fun bnf_of_typ _ _ _ (T as TFree _) accum = ((ID_bnf, ([], [T])), accum)
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   672
  | bnf_of_typ _ _ _ (TVar _) _ = error "Unexpected schematic variable"
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   673
  | bnf_of_typ const_policy qualify' sort (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: 49185
diff changeset
   674
    let
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   675
      val tfrees = Term.add_tfreesT T [];
49236
632f68beff2a full name of a type as key in bnf table
traytel
parents: 49228
diff changeset
   676
      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
   677
    in
49186
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   678
      (case bnf_opt of
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   679
        NONE => ((DEADID_bnf, ([T], [])), (unfold_set, lthy))
49186
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   680
      | SOME bnf =>
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   681
        if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   682
          let
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   683
            val T' = T_of_bnf bnf;
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   684
            val deads = deads_of_bnf bnf;
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   685
            val lives = lives_of_bnf bnf;
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   686
            val tvars' = Term.add_tvarsT T' [];
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   687
            val deads_lives =
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   688
              pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees)))
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   689
                (deads, lives);
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   690
          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: 49185
diff changeset
   691
        else
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   692
          let
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   693
            val name = Long_Name.base_name C;
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   694
            fun qualify i =
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   695
              let val namei = name ^ nonzero_string_of_int i;
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   696
              in qualify' o Binding.qualify true namei end;
49186
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   697
            val odead = dead_of_bnf bnf;
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   698
            val olive = live_of_bnf bnf;
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   699
            val oDs_pos = find_indices [TFree ("dead", [])] (snd (Term.dest_Type
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   700
              (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   701
            val oDs = map (nth Ts) oDs_pos;
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   702
            val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1));
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   703
            val ((inners, (Dss, Ass)), (unfold_set', lthy')) =
49186
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   704
              apfst (apsnd split_list o split_list)
49425
f27f83f71e94 cleaned up internal naming scheme for bnfs
traytel
parents: 49304
diff changeset
   705
                (fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort)
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   706
                (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: 49185
diff changeset
   707
          in
49502
92a7c1842c78 tuned a few ML names
blanchet
parents: 49463
diff changeset
   708
            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: 49185
diff changeset
   709
          end)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   710
    end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   711
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   712
end;