src/HOL/BNF/Tools/bnf_gfp.ML
author blanchet
Fri, 09 Aug 2013 15:40:38 +0200
changeset 52938 3b3e52d5e66b
parent 52923 ec63c82551ae
child 52956 1b62f05ab4fd
permissions -rw-r--r--
tuned name generation code (to make it easier to adapt later)
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_gfp.ML
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
Codatatype construction.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
signature BNF_GFP =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
sig
51867
6d756057e736 signature tuning
blanchet
parents: 51866
diff changeset
    12
  val construct_gfp: mixfix list -> binding list -> binding list -> binding list list ->
6d756057e736 signature tuning
blanchet
parents: 51866
diff changeset
    13
    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
51866
142a82883831 removed dead code
blanchet
parents: 51865
diff changeset
    14
    local_theory -> BNF_FP_Util.fp_result * local_theory
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
structure BNF_GFP : BNF_GFP =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
open BNF_Def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
open BNF_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
open BNF_Tactics
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
    23
open BNF_Comp
51850
106afdf5806c renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
blanchet
parents: 51839
diff changeset
    24
open BNF_FP_Util
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49635
diff changeset
    25
open BNF_FP_Def_Sugar
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
open BNF_GFP_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
open BNF_GFP_Tactics
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    29
datatype wit_tree = Wit_Leaf of int | Wit_Node of (int * int * int list) * wit_tree list;
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    30
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    31
fun mk_tree_args (I, T) (I', Ts) = (sort_distinct int_ord (I @ I'), T :: Ts);
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    32
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    33
fun finish Iss m seen i (nwit, I) =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    34
  let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    35
    val treess = map (fn j =>
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    36
        if j < m orelse member (op =) seen j then [([j], Wit_Leaf j)]
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    37
        else
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    38
          map_index (finish Iss m (insert (op =) j seen) j) (nth Iss (j - m))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    39
          |> flat
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    40
          |> minimize_wits)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    41
      I;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    42
  in
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    43
    map (fn (I, t) => (I, Wit_Node ((i - m, nwit, filter (fn i => i < m) I), t)))
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    44
      (fold_rev (map_product mk_tree_args) treess [([], [])])
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    45
    |> minimize_wits
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    46
  end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    47
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
    48
fun tree_to_ctor_wit vars _ _ (Wit_Leaf j) = ([j], nth vars j)
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
    49
  | tree_to_ctor_wit vars ctors witss (Wit_Node ((i, nwit, I), subtrees)) =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
    50
     (I, nth ctors i $ (Term.list_comb (snd (nth (nth witss i) nwit),
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
    51
       map (snd o tree_to_ctor_wit vars ctors witss) subtrees)));
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    52
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    53
fun tree_to_coind_wits _ (Wit_Leaf _) = []
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    54
  | tree_to_coind_wits lwitss (Wit_Node ((i, nwit, I), subtrees)) =
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    55
     ((i, I), nth (nth lwitss i) nwit) :: maps (tree_to_coind_wits lwitss) subtrees;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
    56
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    57
(*all BNFs have the same lives*)
51867
6d756057e736 signature tuning
blanchet
parents: 51866
diff changeset
    58
fun construct_gfp mixfixes map_bs rel_bs set_bss bs resBs (resDs, Dss) bnfs lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
    val timer = time (Timer.startRealTimer ());
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
    val live = live_of_bnf (hd bnfs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
    val n = length bnfs; (*active*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
    val ks = 1 upto n;
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    65
    val m = live - n (*passive, if 0 don't generate a new BNF*);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
    val ls = 1 upto m;
49498
acc583e14167 tuned variable names
blanchet
parents: 49478
diff changeset
    67
    val b = Binding.name (mk_common_name (map Binding.name_of bs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
    69
    (* TODO: check if m, n, etc., are sane *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
    71
    val deads = fold (union (op =)) Dss resDs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
    val names_lthy = fold Variable.declare_typ deads lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
    (* tvars *)
52938
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    75
    val ((((((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs), passiveXs),
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    76
        passiveYs), idxT) = names_lthy
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    77
      |> mk_TFrees m
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    78
      ||>> mk_TFrees n
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    79
      ||>> mk_TFrees m
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    80
      ||>> mk_TFrees n
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    81
      ||>> mk_TFrees m
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    82
      ||>> mk_TFrees n
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
      ||>> mk_TFrees m
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
      ||>> mk_TFrees m
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
      ||> fst o mk_TFrees 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
      ||> the_single;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
52938
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    88
    val allAs = passiveAs @ activeAs;
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    89
    val allBs' = passiveBs @ activeBs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
    val Ass = replicate n allAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
    val allBs = passiveAs @ activeBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
    val Bss = replicate n allBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
    val allCs = passiveAs @ activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
    val allCs' = passiveBs @ activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
    val Css' = replicate n allCs';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
51866
142a82883831 removed dead code
blanchet
parents: 51865
diff changeset
    97
    (* types *)
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
    98
    val dead_poss =
51866
142a82883831 removed dead code
blanchet
parents: 51865
diff changeset
    99
      map (fn T => if member (op =) deads (TFree T) then SOME (TFree T) else NONE) resBs;
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   100
    fun mk_param NONE passive = (hd passive, tl passive)
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   101
      | mk_param (SOME a) passive = (a, passive);
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   102
    val mk_params = fold_map mk_param dead_poss #> fst;
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   103
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
    fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   105
    val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
49620
7db3d2986881 type of the bound of a BNF depends at most on dead type variables
traytel
parents: 49594
diff changeset
   106
    val (dead_params, dead_params') = `(map Term.dest_TFree) (subtract (op =) passiveAs params');
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
    val FTsAs = mk_FTs allAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
    val FTsBs = mk_FTs allBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
    val FTsCs = mk_FTs allCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
    val ATs = map HOLogic.mk_setT passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
    val BTs = map HOLogic.mk_setT activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
    val B'Ts = map HOLogic.mk_setT activeBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
    val B''Ts = map HOLogic.mk_setT activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
    val sTs = map2 (fn T => fn U => T --> U) activeAs FTsAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
    val s'Ts = map2 (fn T => fn U => T --> U) activeBs FTsBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
    val s''Ts = map2 (fn T => fn U => T --> U) activeCs FTsCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
    val fTs = map2 (fn T => fn U => T --> U) activeAs activeBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
    val self_fTs = map (fn T => T --> T) activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
    val gTs = map2 (fn T => fn U => T --> U) activeBs activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
    val all_gTs = map2 (fn T => fn U => T --> U) allBs allCs';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
    val RTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
    val sRTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
    val R'Ts = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeBs activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
    val setsRTs = map HOLogic.mk_setT sRTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
    val setRTs = map HOLogic.mk_setT RTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
    val all_sbisT = HOLogic.mk_tupleT setsRTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
    val setR'Ts = map HOLogic.mk_setT R'Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
    val FRTs = mk_FTs (passiveAs @ RTs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
    val sumBsAs = map2 (curry mk_sumT) activeBs activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
    val sumFTs = mk_FTs (passiveAs @ sumBsAs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
    val sum_sTs = map2 (fn T => fn U => T --> U) activeAs sumFTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
    (* terms *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
    val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
    val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
    val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
    val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
    val map_Inls = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ sumBsAs)) bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
    val map_Inls_rev = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ sumBsAs)) Bss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
    val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
    val map_snds = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   142
    fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
      (map (replicate live) (replicate n Ts)) bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
    val setssAs = mk_setss allAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
    val setssAs' = transpose setssAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
    val bis_setss = mk_setss (passiveAs @ RTs);
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   147
    val relsAsBs = map4 mk_rel_of_bnf Dss Ass Bss bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   148
    val bds = map3 mk_bd_of_bnf Dss Ass bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   149
    val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
    val sum_bdT = fst (dest_relT (fastype_of sum_bd));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
    val emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   153
    val Zeros = map (fn empty =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
     HOLogic.mk_tuple (map (fn U => absdummy U empty) activeAs)) emptys;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
    val hrecTs = map fastype_of Zeros;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
    val hsetTs = map (fn hrecT => Library.foldr (op -->) (sTs, HOLogic.natT --> hrecT)) hrecTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
   158
    val ((((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2),
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
   159
      z's), As), Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
52298
608afd26a476 killed dead code
blanchet
parents: 52207
diff changeset
   160
      self_fs), gs), all_gs), xFs), yFs), yFs_copy), RFs), (Rtuple, Rtuple')), (hrecs, hrecs')),
608afd26a476 killed dead code
blanchet
parents: 52207
diff changeset
   161
      (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), Kss), names_lthy) = lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
      |> mk_Frees' "b" activeAs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
      ||>> mk_Frees "b" activeAs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
      ||>> mk_Frees "b" activeAs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   165
      ||>> mk_Frees "b" activeBs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
      ||>> mk_Frees "A" ATs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
      ||>> mk_Frees "B" BTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
      ||>> mk_Frees "B" BTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
      ||>> mk_Frees "B'" B'Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
      ||>> mk_Frees "B''" B''Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
      ||>> mk_Frees "s" sTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
      ||>> mk_Frees "sums" sum_sTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
      ||>> mk_Frees "s'" s'Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
      ||>> mk_Frees "s''" s''Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
      ||>> mk_Frees "f" fTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
      ||>> mk_Frees "f" fTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
      ||>> mk_Frees "f" self_fTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
      ||>> mk_Frees "g" gTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
      ||>> mk_Frees "g" all_gTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
      ||>> mk_Frees "x" FTsAs
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   181
      ||>> mk_Frees "y" FTsBs
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   182
      ||>> mk_Frees "y" FTsBs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
      ||>> mk_Frees "x" FRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
      ||>> mk_Frees' "rec" hrecTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
      ||>> mk_Frees "R" setRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
      ||>> mk_Frees "R" setRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
      ||>> mk_Frees "R'" setR'Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
      ||>> mk_Frees "R" setsRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") idxT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
      ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
      ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   194
      ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) activeAs) ATs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   195
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
    val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
   197
    val passive_Id_ons = map mk_Id_on As;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
    val active_UNIVs = map HOLogic.mk_UNIV activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
    val sum_UNIVs = map HOLogic.mk_UNIV sumBsAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
    val passive_ids = map HOLogic.id_const passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
    val active_ids = map HOLogic.id_const activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
    val Inls = map2 Inl_const activeBs activeAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   203
    val fsts = map fst_const RTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
    val snds = map snd_const RTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
    (* thms *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
    val bd_card_orders = map bd_card_order_of_bnf bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
    val bd_card_order = hd bd_card_orders
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
    val bd_Card_orders = map bd_Card_order_of_bnf bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
    val bd_Card_order = hd bd_Card_orders;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   211
    val bd_Cinfinites = map bd_Cinfinite_of_bnf bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
    val bd_Cinfinite = hd bd_Cinfinites;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   213
    val in_monos = map in_mono_of_bnf bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   214
    val map_comps = map map_comp_of_bnf bnfs;
52912
traytel
parents: 52911
diff changeset
   215
    val sym_map_comps = map mk_sym map_comps;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
    val map_comp's = map map_comp'_of_bnf bnfs;
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   217
    val map_cong0s = map map_cong0_of_bnf bnfs;
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
   218
    val map_ids = map map_id_of_bnf bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
    val map_id's = map map_id'_of_bnf bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
    val map_wpulls = map map_wpull_of_bnf bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   221
    val set_bdss = map set_bd_of_bnf bnfs;
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   222
    val set_map'ss = map set_map'_of_bnf bnfs;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   223
    val rel_congs = map rel_cong_of_bnf bnfs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   224
    val rel_converseps = map rel_conversep_of_bnf bnfs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   225
    val rel_Grps = map rel_Grp_of_bnf bnfs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   226
    val rel_eqs = map rel_eq_of_bnf bnfs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   227
    val rel_monos = map rel_mono_of_bnf bnfs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   228
    val rel_OOs = map rel_OO_of_bnf bnfs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   229
    val rel_OO_Grps = map rel_OO_Grp_of_bnf bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
    val timer = time (timer "Extracted terms & thms");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
    (* derived thms *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
    (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x)=
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
      map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
    fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   238
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
        val lhs = Term.list_comb (mapBsCs, all_gs) $
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
          (Term.list_comb (mapAsBs, passive_ids @ fs) $ x);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
        val rhs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
          Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   244
        Goal.prove_sorry lthy [] []
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   245
          (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
          (K (mk_map_comp_id_tac map_comp))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   247
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
      end;
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 map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comp's;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   252
    (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
      map id ... id f(m+1) ... f(m+n) x = x*)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   254
    fun mk_map_cong0L x mapAsAs sets map_cong0 map_id' =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
        fun mk_prem set f z z' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
          HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
            (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
        val prems = map4 mk_prem (drop m sets) self_fs zs zs';
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   260
        val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   262
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
          (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   264
          (K (mk_map_cong0L_tac m map_cong0 map_id'))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   265
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   268
    val map_cong0L_thms = map5 mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_id's;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
    val in_mono'_thms = map (fn thm =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
      (thm OF (replicate m subset_refl)) RS @{thm set_mp}) in_monos;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
    val map_arg_cong_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
      let
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   274
        val prems = map2 (curry mk_Trueprop_eq) yFs yFs_copy;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   275
        val maps = map (fn mapx => Term.list_comb (mapx, all_gs)) mapsBsCs';
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   276
        val concls =
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   277
          map3 (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) yFs yFs_copy maps;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
        val goals =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
          map4 (fn prem => fn concl => fn x => fn y =>
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   280
            fold_rev Logic.all (x :: y :: all_gs) (Logic.mk_implies (prem, concl)))
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   281
          prems concls yFs yFs_copy;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   283
        map (fn goal => Goal.prove_sorry lthy [] [] goal
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
   284
          (K ((hyp_subst_tac lthy THEN' rtac refl) 1)) |> Thm.close_derivation) goals
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
    val timer = time (timer "Derived simple theorems");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
    (* coalgebra *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
    val coalg_bind = Binding.suffix_name ("_" ^ coN ^ algN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
    val coalg_name = Binding.name_of coalg_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
    val coalg_def_bind = (Thm.def_binding coalg_bind, []);
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
    (*forall i = 1 ... n: (\<forall>x \<in> Bi. si \<in> Fi_in A1 .. Am B1 ... Bn)*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
    val coalg_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   297
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   298
        val coalgT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
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 ins = map3 mk_in (replicate n (As @ Bs)) setssAs FTsAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   301
        fun mk_coalg_conjunct B s X z z' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   302
          mk_Ball B (Term.absfree z' (HOLogic.mk_mem (s $ z, X)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
        val lhs = Term.list_comb (Free (coalg_name, coalgT), As @ Bs @ ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
        val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_coalg_conjunct Bs ss ins zs zs')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   307
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   308
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   309
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   310
    val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
   311
      lthy
blanchet
parents: 49308
diff changeset
   312
      |> Specification.definition (SOME (coalg_bind, NONE, NoSyn), (coalg_def_bind, coalg_spec))
blanchet
parents: 49308
diff changeset
   313
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
   314
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   315
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
    val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   317
    val coalg_def = Morphism.thm phi coalg_def_free;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   318
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   319
    fun mk_coalg As Bs ss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   320
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
        val args = As @ Bs @ ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   322
        val Ts = map fastype_of args;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
        val coalgT = Library.foldr (op -->) (Ts, HOLogic.boolT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   325
        Term.list_comb (Const (coalg, coalgT), args)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   328
    val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   329
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
    val coalg_in_thms = map (fn i =>
52904
traytel
parents: 52839
diff changeset
   331
      coalg_def RS iffD1 RS mk_conjunctN n i RS bspec) ks
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   332
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   333
    val coalg_set_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   334
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   335
        val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   336
        fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   337
        fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) B);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   338
        val prems = map2 mk_prem zs Bs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   339
        val conclss = map3 (fn s => fn x => fn sets => map2 (mk_concl s x) (As @ Bs) sets)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   340
          ss zs setssAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   341
        val goalss = map3 (fn x => fn prem => fn concls => map (fn concl =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   342
          fold_rev Logic.all (x :: As @ Bs @ ss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
            (Logic.list_implies (coalg_prem :: [prem], concl))) concls) zs prems conclss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   345
        map (fn goals => map (fn goal => Goal.prove_sorry lthy [] [] goal
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   346
          (K (mk_coalg_set_tac coalg_def)) |> Thm.close_derivation) goals) goalss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   347
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   349
    fun mk_tcoalg ATs BTs = mk_coalg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   350
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   351
    val tcoalg_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   352
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   353
        val goal = fold_rev Logic.all ss
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   354
          (HOLogic.mk_Trueprop (mk_tcoalg passiveAs activeAs ss))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   355
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   356
        Goal.prove_sorry lthy [] [] goal
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
          (K (stac coalg_def 1 THEN CONJ_WRAP
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   358
            (K (EVERY' [rtac ballI, rtac CollectI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   359
              CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   360
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   362
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   363
    val timer = time (timer "Coalgebra definition & thms");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   364
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   365
    (* morphism *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   366
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   367
    val mor_bind = Binding.suffix_name ("_" ^ morN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   368
    val mor_name = Binding.name_of mor_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   369
    val mor_def_bind = (Thm.def_binding mor_bind, []);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   370
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   371
    (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. fi x \<in> B'i)*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   372
    (*mor) forall i = 1 ... n: (\<forall>x \<in> Bi.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   373
       Fi_map id ... id f1 ... fn (si x) = si' (fi x)*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   374
    val mor_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   375
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   376
        val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   377
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   378
        fun mk_fbetw f B1 B2 z z' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
          mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   380
        fun mk_mor B mapAsBs f s s' z z' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   381
          mk_Ball B (Term.absfree z' (HOLogic.mk_eq
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
            (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ z]), s' $ (f $ z))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
        val lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   384
        val rhs = HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   385
          (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   386
           Library.foldr1 HOLogic.mk_conj (map7 mk_mor Bs mapsAsBs fs ss s's zs zs'))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   387
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   388
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   389
      end;
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
    val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
   392
      lthy
blanchet
parents: 49308
diff changeset
   393
      |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
blanchet
parents: 49308
diff changeset
   394
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
   395
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   396
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   397
    val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   398
    val mor_def = Morphism.thm phi mor_def_free;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   399
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   400
    fun mk_mor Bs1 ss1 Bs2 ss2 fs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   401
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   402
        val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   403
        val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   404
        val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   405
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   406
        Term.list_comb (Const (mor, morT), args)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   407
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   408
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
    val mor_prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   411
    val (mor_image_thms, morE_thms) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   412
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   413
        val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   414
        fun mk_image_goal f B1 B2 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   415
          (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   416
        val image_goals = map3 mk_image_goal fs Bs B's;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   417
        fun mk_elim_goal B mapAsBs f s s' x =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
          fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   419
            (Logic.list_implies ([prem, HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B))],
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   420
              mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   421
        val elim_goals = map6 mk_elim_goal Bs mapsAsBs fs ss s's zs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
        fun prove goal =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   423
          Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   424
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   425
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   426
        (map prove image_goals, map prove elim_goals)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   427
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   428
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   429
    val mor_image'_thms = map (fn thm => @{thm set_mp} OF [thm, imageI]) mor_image_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   430
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   431
    val mor_incl_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   433
        val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   434
        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   435
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   436
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   437
          (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   438
          (K (mk_mor_incl_tac mor_def map_id's))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   439
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   440
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   441
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   442
    val mor_id_thm = mor_incl_thm OF (replicate n subset_refl);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   443
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   444
    val mor_comp_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   445
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   446
        val prems =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   447
          [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   448
           HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   449
        val concl =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   450
          HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   451
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   452
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   453
          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   454
            (Logic.list_implies (prems, concl)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   455
          (K (mk_mor_comp_tac mor_def mor_image'_thms morE_thms map_comp_id_thms))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   456
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   457
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   458
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   459
    val mor_cong_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   460
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   461
        val prems = map HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   462
         (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   463
        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   464
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   465
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   466
          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   467
            (Logic.list_implies (prems, concl)))
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
   468
          (K ((hyp_subst_tac lthy THEN' atac) 1))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   469
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   470
      end;
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 mor_UNIV_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   473
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   474
        fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   475
            (HOLogic.mk_comp (Term.list_comb (mapAsBs, passive_ids @ fs), s),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   476
            HOLogic.mk_comp (s', f));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   477
        val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
        val rhs = Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct mapsAsBs fs ss s's);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   479
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   480
        Goal.prove_sorry lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   481
          (K (mk_mor_UNIV_tac morE_thms mor_def))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   482
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   483
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   484
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   485
    val mor_str_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   486
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   487
        val maps = map2 (fn Ds => fn bnf => Term.list_comb
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   488
          (mk_map_of_bnf Ds allAs (passiveAs @ FTsAs) bnf, passive_ids @ ss)) Dss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   489
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   490
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   491
          (fold_rev Logic.all ss (HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   492
            (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   493
          (K (mk_mor_str_tac ks mor_UNIV_thm))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   494
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   495
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   496
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   497
    val mor_sum_case_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   498
      let
49458
blanchet
parents: 49457
diff changeset
   499
        val maps = map3 (fn s => fn sum_s => fn mapx =>
blanchet
parents: 49457
diff changeset
   500
          mk_sum_case (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   501
          s's sum_ss map_Inls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   502
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   503
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   504
          (fold_rev Logic.all (s's @ sum_ss) (HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   505
            (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   506
          (K (mk_mor_sum_case_tac ks mor_UNIV_thm))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   507
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   508
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   509
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   510
    val timer = time (timer "Morphism definition & thms");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   511
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   512
    fun hset_rec_bind j = Binding.suffix_name ("_" ^ hset_recN ^ (if m = 1 then "" else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   513
      string_of_int j)) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   514
    val hset_rec_name = Binding.name_of o hset_rec_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   515
    val hset_rec_def_bind = rpair [] o Thm.def_binding o hset_rec_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   516
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   517
    fun hset_rec_spec j Zero hsetT hrec hrec' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   518
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   519
        fun mk_Suc s setsAs z z' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   520
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   521
            val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setsAs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   522
            fun mk_UN set k = mk_UNION (set $ (s $ z)) (mk_nthN n hrec k);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   523
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   524
            Term.absfree z'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   525
              (mk_union (set $ (s $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   526
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   527
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   528
        val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   529
          (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   530
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   531
        val lhs = Term.list_comb (Free (hset_rec_name j, hsetT), ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   532
        val rhs = mk_nat_rec Zero Suc;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   533
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   534
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   535
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   536
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   537
    val ((hset_rec_frees, (_, hset_rec_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   538
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   539
      |> fold_map5 (fn j => fn Zero => fn hsetT => fn hrec => fn hrec' => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   540
        (SOME (hset_rec_bind j, NONE, NoSyn),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   541
          (hset_rec_def_bind j, hset_rec_spec j Zero hsetT hrec hrec')))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   542
        ls Zeros hsetTs hrecs hrecs'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   543
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   544
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   545
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   546
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   547
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   548
    val hset_rec_defs = map (Morphism.thm phi) hset_rec_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   549
    val hset_recs = map (fst o Term.dest_Const o Morphism.term phi) hset_rec_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   550
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   551
    fun mk_hset_rec ss nat i j T =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   552
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   553
        val args = ss @ [nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   554
        val Ts = map fastype_of ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   555
        val bTs = map domain_type Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   556
        val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) bTs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   557
        val hset_recT = Library.foldr (op -->) (Ts, HOLogic.natT --> hrecT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   558
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   559
        mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   560
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   561
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   562
    val hset_rec_0ss = mk_rec_simps n @{thm nat_rec_0} hset_rec_defs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   563
    val hset_rec_Sucss = mk_rec_simps n @{thm nat_rec_Suc} hset_rec_defs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   564
    val hset_rec_0ss' = transpose hset_rec_0ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   565
    val hset_rec_Sucss' = transpose hset_rec_Sucss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   566
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   567
    fun hset_bind i j = Binding.suffix_name ("_" ^ hsetN ^
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   568
      (if m = 1 then "" else string_of_int j)) (nth bs (i - 1));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   569
    val hset_name = Binding.name_of oo hset_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   570
    val hset_def_bind = rpair [] o Thm.def_binding oo hset_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   571
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   572
    fun hset_spec i j =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   573
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   574
        val U = nth activeAs (i - 1);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   575
        val z = nth zs (i - 1);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   576
        val T = nth passiveAs (j - 1);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   577
        val setT = HOLogic.mk_setT T;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   578
        val hsetT = Library.foldr (op -->) (sTs, U --> setT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   579
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   580
        val lhs = Term.list_comb (Free (hset_name i j, hsetT), ss @ [z]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   581
        val rhs = mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   582
          (Term.absfree nat' (mk_hset_rec ss nat i j T $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   583
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   584
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   585
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   586
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   587
    val ((hset_frees, (_, hset_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   588
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   589
      |> fold_map (fn i => fold_map (fn j => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   590
        (SOME (hset_bind i j, NONE, NoSyn), (hset_def_bind i j, hset_spec i j))) ls) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   591
      |>> map (apsnd split_list o split_list)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   592
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   593
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   594
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   595
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   596
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   597
    val hset_defss = map (map (Morphism.thm phi)) hset_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   598
    val hset_defss' = transpose hset_defss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   599
    val hset_namess = map (map (fst o Term.dest_Const o Morphism.term phi)) hset_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   600
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   601
    fun mk_hset ss i j T =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   602
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   603
        val Ts = map fastype_of ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   604
        val bTs = map domain_type Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   605
        val hsetT = Library.foldr (op -->) (Ts, nth bTs (i - 1) --> HOLogic.mk_setT T);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   606
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   607
        Term.list_comb (Const (nth (nth hset_namess (i - 1)) (j - 1), hsetT), ss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   608
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   609
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   610
    val hsetssAs = map (fn i => map2 (mk_hset ss i) ls passiveAs) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   611
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   612
    val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   613
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   614
        fun mk_set_incl_hset s x set hset = fold_rev Logic.all (x :: ss)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   615
          (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (hset $ x)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   616
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   617
        fun mk_set_hset_incl_hset s x y set hset1 hset2 =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   618
          fold_rev Logic.all (x :: y :: ss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   619
            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (s $ y))),
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   620
            HOLogic.mk_Trueprop (mk_leq (hset1 $ x) (hset2 $ y))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   621
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   622
        val set_incl_hset_goalss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   623
          map4 (fn s => fn x => fn sets => fn hsets =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   624
            map2 (mk_set_incl_hset s x) (take m sets) hsets)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   625
          ss zs setssAs hsetssAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   626
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   627
        (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   628
        val set_hset_incl_hset_goalsss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   629
          map4 (fn si => fn yi => fn sets => fn hsetsi =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   630
            map3 (fn xk => fn set => fn hsetsk =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   631
              map2 (mk_set_hset_incl_hset si xk yi set) hsetsk hsetsi)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   632
            zs_copy (drop m sets) hsetssAs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   633
          ss zs setssAs hsetssAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   634
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   635
        (map3 (fn goals => fn defs => fn rec_Sucs =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   636
          map3 (fn goal => fn def => fn rec_Suc =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   637
            Goal.prove_sorry lthy [] [] goal (K (mk_set_incl_hset_tac def rec_Suc))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   638
            |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   639
          goals defs rec_Sucs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   640
        set_incl_hset_goalss hset_defss hset_rec_Sucss,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   641
        map3 (fn goalss => fn defsi => fn rec_Sucs =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   642
          map3 (fn k => fn goals => fn defsk =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   643
            map4 (fn goal => fn defk => fn defi => fn rec_Suc =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   644
              Goal.prove_sorry lthy [] [] goal
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   645
                (K (mk_set_hset_incl_hset_tac n [defk, defi] rec_Suc k))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   646
              |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   647
            goals defsk defsi rec_Sucs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   648
          ks goalss hset_defss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   649
        set_hset_incl_hset_goalsss hset_defss hset_rec_Sucss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   650
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   651
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   652
    val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   653
    val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   654
    val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   655
    val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   656
      set_hset_incl_hset_thmsss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   657
    val set_hset_thmss' = transpose set_hset_thmss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   658
    val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   659
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   660
    val set_incl_hin_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   661
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   662
        fun mk_set_incl_hin s x hsets1 set hsets2 T =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   663
          fold_rev Logic.all (x :: ss @ As)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   664
            (Logic.list_implies
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   665
              (map2 (fn hset => fn A => HOLogic.mk_Trueprop (mk_leq (hset $ x) A)) hsets1 As,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   666
              HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (mk_in As hsets2 T))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   667
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   668
        val set_incl_hin_goalss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   669
          map4 (fn s => fn x => fn sets => fn hsets =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   670
            map3 (mk_set_incl_hin s x hsets) (drop m sets) hsetssAs activeAs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   671
          ss zs setssAs hsetssAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   672
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   673
        map2 (map2 (fn goal => fn thms =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   674
          Goal.prove_sorry lthy [] [] goal (K (mk_set_incl_hin_tac thms))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   675
          |> Thm.close_derivation))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   676
        set_incl_hin_goalss set_hset_incl_hset_thmsss
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   677
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   678
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   679
    val hset_minimal_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   680
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   681
        fun mk_passive_prem set s x K =
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   682
          Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (K $ x)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   683
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   684
        fun mk_active_prem s x1 K1 set x2 K2 =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   685
          fold_rev Logic.all [x1, x2]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   686
            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (s $ x1))),
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   687
              HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   688
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   689
        val premss = map2 (fn j => fn Ks =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   690
          map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) setssAs) ss zs Ks @
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   691
            flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   692
              map3 (mk_active_prem s x1 K1) (drop m sets) zs_copy Ks) setssAs ss zs Ks))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   693
          ls Kss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   694
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   695
        val hset_rec_minimal_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   696
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   697
            fun mk_conjunct j T i K x = mk_leq (mk_hset_rec ss nat i j T $ x) (K $ x);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   698
            fun mk_concl j T Ks = list_all_free zs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   699
              (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   700
            val concls = map3 mk_concl ls passiveAs Kss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   701
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   702
            val goals = map2 (fn prems => fn concl =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   703
              Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   704
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   705
            val ctss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   706
              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   707
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   708
            map4 (fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   709
              singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   710
                (Goal.prove_sorry lthy [] [] goal
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   711
                  (mk_hset_rec_minimal_tac m cts hset_rec_0s hset_rec_Sucs))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   712
              |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   713
            goals ctss hset_rec_0ss' hset_rec_Sucss'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   714
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   715
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   716
        fun mk_conjunct j T i K x = mk_leq (mk_hset ss i j T $ x) (K $ x);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   717
        fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   718
        val concls = map3 mk_concl ls passiveAs Kss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   719
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   720
        val goals = map3 (fn Ks => fn prems => fn concl =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   721
          fold_rev Logic.all (Ks @ ss @ zs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   722
            (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   723
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   724
        map3 (fn goal => fn hset_defs => fn hset_rec_minimal =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   725
          Goal.prove_sorry lthy [] [] goal
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   726
            (mk_hset_minimal_tac n hset_defs hset_rec_minimal)
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   727
          |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   728
        goals hset_defss' hset_rec_minimal_thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   729
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   730
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   731
    val timer = time (timer "Hereditary sets");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   732
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   733
    (* bisimulation *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   734
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   735
    val bis_bind = Binding.suffix_name ("_" ^ bisN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   736
    val bis_name = Binding.name_of bis_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   737
    val bis_def_bind = (Thm.def_binding bis_bind, []);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   738
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   739
    fun mk_bis_le_conjunct R B1 B2 = mk_leq R (mk_Times (B1, B2));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   740
    val bis_le = Library.foldr1 HOLogic.mk_conj (map3 mk_bis_le_conjunct Rs Bs B's)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   741
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   742
    val bis_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   743
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   744
        val bisT = Library.foldr (op -->) (ATs @ BTs @ sTs @ B'Ts @ s'Ts @ setRTs, HOLogic.boolT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   745
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   746
        val fst_args = passive_ids @ fsts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   747
        val snd_args = passive_ids @ snds;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   748
        fun mk_bis R s s' b1 b2 RF map1 map2 sets =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   749
          list_all_free [b1, b2] (HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   750
            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   751
            mk_Bex (mk_in (As @ Rs) sets (snd (dest_Free RF))) (Term.absfree (dest_Free RF)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   752
              (HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   753
                (HOLogic.mk_eq (Term.list_comb (map1, fst_args) $ RF, s $ b1),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   754
                HOLogic.mk_eq (Term.list_comb (map2, snd_args) $ RF, s' $ b2))))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   755
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   756
        val lhs = Term.list_comb (Free (bis_name, bisT), As @ Bs @ ss @ B's @ s's @ Rs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   757
        val rhs = HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   758
          (bis_le, Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   759
            (map9 mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   760
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   761
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   762
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   763
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   764
    val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
   765
      lthy
blanchet
parents: 49308
diff changeset
   766
      |> Specification.definition (SOME (bis_bind, NONE, NoSyn), (bis_def_bind, bis_spec))
blanchet
parents: 49308
diff changeset
   767
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
   768
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   769
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   770
    val bis = fst (Term.dest_Const (Morphism.term phi bis_free));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   771
    val bis_def = Morphism.thm phi bis_def_free;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   772
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   773
    fun mk_bis As Bs1 ss1 Bs2 ss2 Rs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   774
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   775
        val args = As @ Bs1 @ ss1 @ Bs2 @ ss2 @ Rs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   776
        val Ts = map fastype_of args;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   777
        val bisT = Library.foldr (op -->) (Ts, HOLogic.boolT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   778
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   779
        Term.list_comb (Const (bis, bisT), args)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   780
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   781
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   782
    val bis_cong_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   783
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   784
        val prems = map HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   785
         (mk_bis As Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   786
        val concl = HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs_copy);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   787
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   788
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   789
          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs @ Rs_copy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   790
            (Logic.list_implies (prems, concl)))
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
   791
          (K ((hyp_subst_tac lthy THEN' atac) 1))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   792
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   793
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   794
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   795
    val bis_rel_thm =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   796
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   797
        fun mk_conjunct R s s' b1 b2 rel =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   798
          list_all_free [b1, b2] (HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   799
            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   800
            Term.list_comb (rel, map mk_in_rel (passive_Id_ons @ Rs)) $ (s $ b1) $ (s' $ b2)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   801
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   802
        val rhs = HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   803
          (bis_le, Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   804
            (map6 mk_conjunct Rs ss s's zs z's relsAsBs))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   805
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   806
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   807
          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   808
            (mk_Trueprop_eq (mk_bis As Bs ss B's s's Rs, rhs)))
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   809
          (K (mk_bis_rel_tac lthy m bis_def rel_OO_Grps map_comp's map_cong0s set_map'ss))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   810
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   811
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   812
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   813
    val bis_converse_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   814
      Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   815
        (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   816
          (Logic.mk_implies
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   817
            (HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   818
            HOLogic.mk_Trueprop (mk_bis As B's s's Bs ss (map mk_converse Rs)))))
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   819
        (K (mk_bis_converse_tac m bis_rel_thm rel_congs rel_converseps))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   820
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   821
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   822
    val bis_O_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   823
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   824
        val prems =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   825
          [HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   826
           HOLogic.mk_Trueprop (mk_bis As B's s's B''s s''s R's)];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   827
        val concl =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   828
          HOLogic.mk_Trueprop (mk_bis As Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   829
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   830
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   831
          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ B''s @ s''s @ Rs @ R's)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   832
            (Logic.list_implies (prems, concl)))
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   833
          (K (mk_bis_O_tac lthy m bis_rel_thm rel_congs rel_OOs))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   834
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   835
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   836
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   837
    val bis_Gr_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   838
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   839
        val concl =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   840
          HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map2 mk_Gr Bs fs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   841
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   842
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   843
          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ fs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   844
            (Logic.list_implies ([coalg_prem, mor_prem], concl)))
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   845
          (mk_bis_Gr_tac bis_rel_thm rel_Grps mor_image_thms morE_thms coalg_in_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   846
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   847
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   848
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   849
    val bis_image2_thm = bis_cong_thm OF
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   850
      ((bis_O_thm OF [bis_Gr_thm RS bis_converse_thm, bis_Gr_thm]) ::
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   851
      replicate n @{thm image2_Gr});
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   852
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
   853
    val bis_Id_on_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) ::
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
   854
      replicate n @{thm Id_on_Gr});
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   855
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   856
    val bis_Union_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   857
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   858
        val prem =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   859
          HOLogic.mk_Trueprop (mk_Ball Idx
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   860
            (Term.absfree idx' (mk_bis As Bs ss B's s's (map (fn R => R $ idx) Ris))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   861
        val concl =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   862
          HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map (mk_UNION Idx) Ris));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   863
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   864
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   865
          (fold_rev Logic.all (Idx :: As @ Bs @ ss @ B's @ s's @ Ris)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   866
            (Logic.mk_implies (prem, concl)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   867
          (mk_bis_Union_tac bis_def in_mono'_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   868
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   869
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   870
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   871
    (* self-bisimulation *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   872
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   873
    fun mk_sbis As Bs ss Rs = mk_bis As Bs ss Bs ss Rs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   874
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   875
    val sbis_prem = HOLogic.mk_Trueprop (mk_sbis As Bs ss sRs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   876
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   877
    (* largest self-bisimulation *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   878
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   879
    fun lsbis_bind i = Binding.suffix_name ("_" ^ lsbisN ^ (if n = 1 then "" else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   880
      string_of_int i)) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   881
    val lsbis_name = Binding.name_of o lsbis_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   882
    val lsbis_def_bind = rpair [] o Thm.def_binding o lsbis_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   883
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   884
    val all_sbis = HOLogic.mk_Collect (fst Rtuple', snd Rtuple', list_exists_free sRs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   885
      (HOLogic.mk_conj (HOLogic.mk_eq (Rtuple, HOLogic.mk_tuple sRs), mk_sbis As Bs ss sRs)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   886
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   887
    fun lsbis_spec i RT =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   888
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   889
        fun mk_lsbisT RT =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   890
          Library.foldr (op -->) (map fastype_of (As @ Bs @ ss), RT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   891
        val lhs = Term.list_comb (Free (lsbis_name i, mk_lsbisT RT), As @ Bs @ ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   892
        val rhs = mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   893
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   894
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   895
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   896
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   897
    val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   898
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   899
      |> fold_map2 (fn i => fn RT => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   900
        (SOME (lsbis_bind i, NONE, NoSyn), (lsbis_def_bind i, lsbis_spec i RT))) ks setsRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   901
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   902
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   903
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   904
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   905
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   906
    val lsbis_defs = map (Morphism.thm phi) lsbis_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   907
    val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   908
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   909
    fun mk_lsbis As Bs ss i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   910
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   911
        val args = As @ Bs @ ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   912
        val Ts = map fastype_of args;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   913
        val RT = mk_relT (`I (HOLogic.dest_setT (fastype_of (nth Bs (i - 1)))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   914
        val lsbisT = Library.foldr (op -->) (Ts, RT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   915
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   916
        Term.list_comb (Const (nth lsbiss (i - 1), lsbisT), args)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   917
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   918
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   919
    val sbis_lsbis_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   920
      Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   921
        (fold_rev Logic.all (As @ Bs @ ss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   922
          (HOLogic.mk_Trueprop (mk_sbis As Bs ss (map (mk_lsbis As Bs ss) ks))))
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
   923
        (K (mk_sbis_lsbis_tac lthy lsbis_defs bis_Union_thm bis_cong_thm))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   924
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   925
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   926
    val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS
52904
traytel
parents: 52839
diff changeset
   927
      (bis_def RS iffD1 RS conjunct1 RS mk_conjunctN n i)) ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   928
    val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS
52904
traytel
parents: 52839
diff changeset
   929
      (bis_def RS iffD1 RS conjunct2 RS mk_conjunctN n i))) RS mp) ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   930
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   931
    val incl_lsbis_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   932
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   933
        fun mk_concl i R = HOLogic.mk_Trueprop (mk_leq R (mk_lsbis As Bs ss i));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   934
        val goals = map2 (fn i => fn R => fold_rev Logic.all (As @ Bs @ ss @ sRs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   935
          (Logic.mk_implies (sbis_prem, mk_concl i R))) ks sRs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   936
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   937
        map3 (fn goal => fn i => fn def => Goal.prove_sorry lthy [] [] goal
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   938
          (K (mk_incl_lsbis_tac n i def)) |> Thm.close_derivation) goals ks lsbis_defs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   939
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   940
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   941
    val equiv_lsbis_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   942
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   943
        fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis As Bs ss i));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   944
        val goals = map2 (fn i => fn B => fold_rev Logic.all (As @ Bs @ ss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   945
          (Logic.mk_implies (coalg_prem, mk_concl i B))) ks Bs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   946
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   947
        map3 (fn goal => fn l_incl => fn incl_l =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   948
          Goal.prove_sorry lthy [] [] goal
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   949
            (K (mk_equiv_lsbis_tac sbis_lsbis_thm l_incl incl_l
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
   950
              bis_Id_on_thm bis_converse_thm bis_O_thm))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   951
          |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   952
        goals lsbis_incl_thms incl_lsbis_thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   953
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   954
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   955
    val timer = time (timer "Bisimulations");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   956
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   957
    (* bounds *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   958
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   959
    val (lthy, sbd, sbdT,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
   960
      sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   961
      if n = 1
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
   962
      then (lthy, sum_bd, sum_bdT, bd_card_order, bd_Cinfinite, bd_Card_order, set_bdss)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   963
      else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   964
        let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   965
          val sbdT_bind = Binding.suffix_name ("_" ^ sum_bdTN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   966
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   967
          val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =
49835
31f32ec4d766 discontinued typedef with alternative name;
wenzelm
parents: 49833
diff changeset
   968
            typedef (sbdT_bind, dead_params, NoSyn)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   969
              (HOLogic.mk_UNIV sum_bdT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   970
49620
7db3d2986881 type of the bound of a BNF depends at most on dead type variables
traytel
parents: 49594
diff changeset
   971
          val sbdT = Type (sbdT_name, dead_params');
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   972
          val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   973
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   974
          val sbd_bind = Binding.suffix_name ("_" ^ sum_bdN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   975
          val sbd_name = Binding.name_of sbd_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   976
          val sbd_def_bind = (Thm.def_binding sbd_bind, []);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   977
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   978
          val sbd_spec = HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   979
            (HOLogic.mk_eq (Free (sbd_name, mk_relT (`I sbdT)), mk_dir_image sum_bd Abs_sbdT));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   980
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   981
          val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   982
            lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   983
            |> Specification.definition (SOME (sbd_bind, NONE, NoSyn), (sbd_def_bind, sbd_spec))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   984
            ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   985
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   986
          val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   987
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   988
          val sbd_def = Morphism.thm phi sbd_def_free;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   989
          val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   990
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49225
diff changeset
   991
          val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info);
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49225
diff changeset
   992
          val Abs_sbdT_bij = mk_Abs_bij_thm lthy Abs_sbdT_inj (#Abs_cases sbdT_loc_info);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   993
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   994
          fun mk_sum_Cinfinite [thm] = thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   995
            | mk_sum_Cinfinite (thm :: thms) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   996
              @{thm Cinfinite_csum_strong} OF [thm, mk_sum_Cinfinite thms];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   997
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   998
          val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   999
          val sum_Card_order = sum_Cinfinite RS conjunct2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1000
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1001
          fun mk_sum_card_order [thm] = thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1002
            | mk_sum_card_order (thm :: thms) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1003
              @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1004
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1005
          val sum_card_order = mk_sum_card_order bd_card_orders;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1006
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1007
          val sbd_ordIso = fold_thms lthy [sbd_def]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1008
            (@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order]);
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1009
          val sbd_card_order =  fold_thms lthy [sbd_def]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1010
            (@{thm card_order_dir_image} OF [Abs_sbdT_bij, sum_card_order]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1011
          val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1012
          val sbd_Card_order = sbd_Cinfinite RS conjunct2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1013
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1014
          fun mk_set_sbd i bd_Card_order bds =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1015
            map (fn thm => @{thm ordLeq_ordIso_trans} OF
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1016
              [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1017
          val set_sbdss = map3 mk_set_sbd ks bd_Card_orders set_bdss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1018
       in
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
  1019
         (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1020
       end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1021
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1022
    val sbdTs = replicate n sbdT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1023
    val sum_sbd = Library.foldr1 (uncurry mk_csum) (replicate n sbd);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1024
    val sum_sbdT = mk_sumTN sbdTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1025
    val sum_sbd_listT = HOLogic.listT sum_sbdT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1026
    val sum_sbd_list_setT = HOLogic.mk_setT sum_sbd_listT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1027
    val bdTs = passiveAs @ replicate n sbdT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1028
    val to_sbd_maps = map4 mk_map_of_bnf Dss Ass (replicate n bdTs) bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1029
    val bdFTs = mk_FTs bdTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1030
    val sbdFT = mk_sumTN bdFTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1031
    val treeT = HOLogic.mk_prodT (sum_sbd_list_setT, sum_sbd_listT --> sbdFT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1032
    val treeQT = HOLogic.mk_setT treeT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1033
    val treeTs = passiveAs @ replicate n treeT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1034
    val treeQTs = passiveAs @ replicate n treeQT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1035
    val treeFTs = mk_FTs treeTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1036
    val tree_maps = map4 mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1037
    val final_maps = map4 mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1038
    val isNode_setss = mk_setss (passiveAs @ replicate n sbdT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1039
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1040
    val root = HOLogic.mk_set sum_sbd_listT [HOLogic.mk_list sum_sbdT []];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1041
    val Zero = HOLogic.mk_tuple (map (fn U => absdummy U root) activeAs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1042
    val Lev_recT = fastype_of Zero;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1043
    val LevT = Library.foldr (op -->) (sTs, HOLogic.natT --> Lev_recT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1044
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1045
    val Nil = HOLogic.mk_tuple (map3 (fn i => fn z => fn z'=>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1046
      Term.absfree z' (mk_InN activeAs z i)) ks zs zs');
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1047
    val rv_recT = fastype_of Nil;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1048
    val rvT = Library.foldr (op -->) (sTs, sum_sbd_listT --> rv_recT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1049
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1050
    val (((((((((((sumx, sumx'), (kks, kks')), (kl, kl')), (kl_copy, kl'_copy)), (Kl, Kl')),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1051
      (lab, lab')), (Kl_lab, Kl_lab')), xs), (Lev_rec, Lev_rec')), (rv_rec, rv_rec')),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1052
      names_lthy) = names_lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1053
      |> yield_singleton (apfst (op ~~) oo mk_Frees' "sumx") sum_sbdT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1054
      ||>> mk_Frees' "k" sbdTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1055
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1056
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1057
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl") sum_sbd_list_setT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1058
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "lab") (sum_sbd_listT --> sbdFT)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1059
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl_lab") treeT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1060
      ||>> mk_Frees "x" bdFTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1061
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") Lev_recT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1062
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") rv_recT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1063
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1064
    val (k, k') = (hd kks, hd kks')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1065
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1066
    val timer = time (timer "Bounds");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1067
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1068
    (* tree coalgebra *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1069
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1070
    fun isNode_bind i = Binding.suffix_name ("_" ^ isNodeN ^ (if n = 1 then "" else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1071
      string_of_int i)) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1072
    val isNode_name = Binding.name_of o isNode_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1073
    val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1074
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1075
    val isNodeT =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1076
      Library.foldr (op -->) (map fastype_of (As @ [Kl, lab, kl]), HOLogic.boolT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1077
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1078
    val Succs = map3 (fn i => fn k => fn k' =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1079
      HOLogic.mk_Collect (fst k', snd k', HOLogic.mk_mem (mk_InN sbdTs k i, mk_Succ Kl kl)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1080
      ks kks kks';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1081
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1082
    fun isNode_spec sets x i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1083
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1084
        val (passive_sets, active_sets) = chop m (map (fn set => set $ x) sets);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1085
        val lhs = Term.list_comb (Free (isNode_name i, isNodeT), As @ [Kl, lab, kl]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1086
        val rhs = list_exists_free [x]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1087
          (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) ::
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1088
          map2 mk_leq passive_sets As @ map2 (curry HOLogic.mk_eq) active_sets Succs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1089
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1090
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1091
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1092
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1093
    val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1094
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1095
      |> fold_map3 (fn i => fn x => fn sets => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1096
        (SOME (isNode_bind i, NONE, NoSyn), (isNode_def_bind i, isNode_spec sets x i)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1097
        ks xs isNode_setss
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1098
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1099
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1100
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1101
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1102
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1103
    val isNode_defs = map (Morphism.thm phi) isNode_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1104
    val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1105
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1106
    fun mk_isNode As kl i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1107
      Term.list_comb (Const (nth isNodes (i - 1), isNodeT), As @ [Kl, lab, kl]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1108
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1109
    val isTree =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1110
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1111
        val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl);
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1112
        val Field = mk_leq Kl (mk_Field (mk_clists sum_sbd));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1113
        val prefCl = mk_prefCl Kl;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1114
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1115
        val tree = mk_Ball Kl (Term.absfree kl'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1116
          (HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1117
            (Library.foldr1 HOLogic.mk_disj (map (mk_isNode As kl) ks),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1118
            Library.foldr1 HOLogic.mk_conj (map4 (fn Succ => fn i => fn k => fn k' =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1119
              mk_Ball Succ (Term.absfree k' (mk_isNode As
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1120
                (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1121
            Succs ks kks kks'))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1122
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1123
        val undef = list_all_free [kl] (HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1124
          (HOLogic.mk_not (HOLogic.mk_mem (kl, Kl)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1125
          HOLogic.mk_eq (lab $ kl, mk_undefined sbdFT)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1126
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1127
        Library.foldr1 HOLogic.mk_conj [empty, Field, prefCl, tree, undef]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1128
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1129
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1130
    fun carT_bind i = Binding.suffix_name ("_" ^ carTN ^ (if n = 1 then "" else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1131
      string_of_int i)) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1132
    val carT_name = Binding.name_of o carT_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1133
    val carT_def_bind = rpair [] o Thm.def_binding o carT_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1134
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1135
    fun carT_spec i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1136
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1137
        val carTT = Library.foldr (op -->) (ATs, HOLogic.mk_setT treeT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1138
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1139
        val lhs = Term.list_comb (Free (carT_name i, carTT), As);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1140
        val rhs = HOLogic.mk_Collect (fst Kl_lab', snd Kl_lab', list_exists_free [Kl, lab]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1141
          (HOLogic.mk_conj (HOLogic.mk_eq (Kl_lab, HOLogic.mk_prod (Kl, lab)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1142
            HOLogic.mk_conj (isTree, mk_isNode As (HOLogic.mk_list sum_sbdT []) i))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1143
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1144
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1145
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1146
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1147
    val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1148
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1149
      |> fold_map (fn i => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1150
        (SOME (carT_bind i, NONE, NoSyn), (carT_def_bind i, carT_spec i))) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1151
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1152
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1153
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1154
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1155
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1156
    val carT_defs = map (Morphism.thm phi) carT_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1157
    val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1158
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1159
    fun mk_carT As i = Term.list_comb
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1160
      (Const (nth carTs (i - 1),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1161
         Library.foldr (op -->) (map fastype_of As, HOLogic.mk_setT treeT)), As);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1162
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1163
    fun strT_bind i = Binding.suffix_name ("_" ^ strTN ^ (if n = 1 then "" else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1164
      string_of_int i)) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1165
    val strT_name = Binding.name_of o strT_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1166
    val strT_def_bind = rpair [] o Thm.def_binding o strT_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1167
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1168
    fun strT_spec mapFT FT i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1169
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1170
        val strTT = treeT --> FT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1171
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1172
        fun mk_f i k k' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1173
          let val in_k = mk_InN sbdTs k i;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1174
          in Term.absfree k' (HOLogic.mk_prod (mk_Shift Kl in_k, mk_shift lab in_k)) end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1175
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1176
        val f = Term.list_comb (mapFT, passive_ids @ map3 mk_f ks kks kks');
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1177
        val (fTs1, fTs2) = apsnd tl (chop (i - 1) (map (fn T => T --> FT) bdFTs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1178
        val fs = map mk_undefined fTs1 @ (f :: map mk_undefined fTs2);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1179
        val lhs = Free (strT_name i, strTT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1180
        val rhs = HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1181
          (mk_sum_caseN fs $ (lab $ HOLogic.mk_list sum_sbdT []))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1182
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1183
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1184
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1185
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1186
    val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1187
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1188
      |> fold_map3 (fn i => fn mapFT => fn FT => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1189
        (SOME (strT_bind i, NONE, NoSyn), (strT_def_bind i, strT_spec mapFT FT i)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1190
        ks tree_maps treeFTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1191
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1192
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1193
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1194
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1195
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1196
    val strT_defs = map ((fn def => trans OF [def RS fun_cong, @{thm prod.cases}]) o
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1197
      Morphism.thm phi) strT_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1198
    val strTs = map (fst o Term.dest_Const o Morphism.term phi) strT_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1199
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1200
    fun mk_strT FT i = Const (nth strTs (i - 1), treeT --> FT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1201
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1202
    val carTAs = map (mk_carT As) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1203
    val strTAs = map2 mk_strT treeFTs ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1204
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1205
    val coalgT_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1206
      Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1207
        (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_coalg As carTAs strTAs)))
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  1208
        (mk_coalgT_tac m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_map'ss)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1209
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1210
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1211
    val timer = time (timer "Tree coalgebra");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1212
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1213
    fun mk_to_sbd s x i i' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1214
      mk_toCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1215
    fun mk_from_sbd s x i i' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1216
      mk_fromCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1217
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1218
    fun mk_to_sbd_thmss thm = map (map (fn set_sbd =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1219
      thm OF [set_sbd, sbd_Card_order]) o drop m) set_sbdss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1220
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1221
    val to_sbd_inj_thmss = mk_to_sbd_thmss @{thm toCard_inj};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1222
    val to_sbd_thmss = mk_to_sbd_thmss @{thm toCard};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1223
    val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1224
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1225
    val Lev_bind = Binding.suffix_name ("_" ^ LevN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1226
    val Lev_name = Binding.name_of Lev_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1227
    val Lev_def_bind = rpair [] (Thm.def_binding Lev_bind);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1228
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1229
    val Lev_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1230
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1231
        fun mk_Suc i s setsAs a a' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1232
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1233
            val sets = drop m setsAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1234
            fun mk_set i' set b =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1235
              let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1236
                val Cons = HOLogic.mk_eq (kl_copy,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1237
                  mk_Cons (mk_InN sbdTs (mk_to_sbd s a i i' $ b) i') kl)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1238
                val b_set = HOLogic.mk_mem (b, set $ (s $ a));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1239
                val kl_rec = HOLogic.mk_mem (kl, mk_nthN n Lev_rec i' $ b);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1240
              in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1241
                HOLogic.mk_Collect (fst kl'_copy, snd kl'_copy, list_exists_free [b, kl]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1242
                  (HOLogic.mk_conj (Cons, HOLogic.mk_conj (b_set, kl_rec))))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1243
              end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1244
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1245
            Term.absfree a' (Library.foldl1 mk_union (map3 mk_set ks sets zs_copy))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1246
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1247
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1248
        val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1249
          (HOLogic.mk_tuple (map5 mk_Suc ks ss setssAs zs zs')));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1250
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1251
        val lhs = Term.list_comb (Free (Lev_name, LevT), ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1252
        val rhs = mk_nat_rec Zero Suc;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1253
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1254
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1255
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1256
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1257
    val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1258
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1259
      |> Specification.definition (SOME (Lev_bind, NONE, NoSyn), (Lev_def_bind, Lev_spec))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1260
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1261
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1262
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1263
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1264
    val Lev_def = Morphism.thm phi Lev_def_free;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1265
    val Lev = fst (Term.dest_Const (Morphism.term phi Lev_free));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1266
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1267
    fun mk_Lev ss nat i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1268
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1269
        val Ts = map fastype_of ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1270
        val LevT = Library.foldr (op -->) (Ts, HOLogic.natT -->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1271
          HOLogic.mk_tupleT (map (fn U => domain_type U --> sum_sbd_list_setT) Ts));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1272
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1273
        mk_nthN n (Term.list_comb (Const (Lev, LevT), ss) $ nat) i
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1274
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1275
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1276
    val Lev_0s = flat (mk_rec_simps n @{thm nat_rec_0} [Lev_def]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1277
    val Lev_Sucs = flat (mk_rec_simps n @{thm nat_rec_Suc} [Lev_def]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1278
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1279
    val rv_bind = Binding.suffix_name ("_" ^ rvN) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1280
    val rv_name = Binding.name_of rv_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1281
    val rv_def_bind = rpair [] (Thm.def_binding rv_bind);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1282
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1283
    val rv_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1284
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1285
        fun mk_Cons i s b b' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1286
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1287
            fun mk_case i' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1288
              Term.absfree k' (mk_nthN n rv_rec i' $ (mk_from_sbd s b i i' $ k));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1289
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1290
            Term.absfree b' (mk_sum_caseN (map mk_case ks) $ sumx)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1291
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1292
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1293
        val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1294
          (HOLogic.mk_tuple (map4 mk_Cons ks ss zs zs'))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1295
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1296
        val lhs = Term.list_comb (Free (rv_name, rvT), ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1297
        val rhs = mk_list_rec Nil Cons;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1298
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1299
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1300
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1301
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1302
    val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1303
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1304
      |> Specification.definition (SOME (rv_bind, NONE, NoSyn), (rv_def_bind, rv_spec))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1305
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1306
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1307
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1308
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1309
    val rv_def = Morphism.thm phi rv_def_free;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1310
    val rv = fst (Term.dest_Const (Morphism.term phi rv_free));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1311
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1312
    fun mk_rv ss kl i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1313
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1314
        val Ts = map fastype_of ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1315
        val As = map domain_type Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1316
        val rvT = Library.foldr (op -->) (Ts, fastype_of kl -->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1317
          HOLogic.mk_tupleT (map (fn U => U --> mk_sumTN As) As));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1318
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1319
        mk_nthN n (Term.list_comb (Const (rv, rvT), ss) $ kl) i
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1320
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1321
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1322
    val rv_Nils = flat (mk_rec_simps n @{thm list_rec_Nil} [rv_def]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1323
    val rv_Conss = flat (mk_rec_simps n @{thm list_rec_Cons} [rv_def]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1324
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1325
    fun beh_bind i = Binding.suffix_name ("_" ^ behN ^ (if n = 1 then "" else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1326
      string_of_int i)) b;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1327
    val beh_name = Binding.name_of o beh_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1328
    val beh_def_bind = rpair [] o Thm.def_binding o beh_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1329
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1330
    fun beh_spec i z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1331
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1332
        val mk_behT = Library.foldr (op -->) (map fastype_of (ss @ [z]), treeT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1333
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1334
        fun mk_case i to_sbd_map s k k' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1335
          Term.absfree k' (mk_InN bdFTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1336
            (Term.list_comb (to_sbd_map, passive_ids @ map (mk_to_sbd s k i) ks) $ (s $ k)) i);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1337
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1338
        val Lab = Term.absfree kl' (mk_If
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1339
          (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1340
          (mk_sum_caseN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1341
          (mk_undefined sbdFT));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1342
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1343
        val lhs = Term.list_comb (Free (beh_name i, mk_behT), ss) $ z;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1344
        val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1345
          (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1346
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1347
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1348
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1349
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1350
    val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1351
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1352
      |> fold_map2 (fn i => fn z => Specification.definition
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1353
        (SOME (beh_bind i, NONE, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1354
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1355
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1356
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1357
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1358
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1359
    val beh_defs = map (Morphism.thm phi) beh_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1360
    val behs = map (fst o Term.dest_Const o Morphism.term phi) beh_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1361
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1362
    fun mk_beh ss i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1363
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1364
        val Ts = map fastype_of ss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1365
        val behT = Library.foldr (op -->) (Ts, nth activeAs (i - 1) --> treeT);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1366
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1367
        Term.list_comb (Const (nth behs (i - 1), behT), ss)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1368
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1369
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1370
    val Lev_sbd_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1371
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1372
        fun mk_conjunct i z = mk_leq (mk_Lev ss nat i $ z) (mk_Field (mk_clists sum_sbd));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1373
        val goal = list_all_free zs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1374
          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1375
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1376
        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1377
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1378
        val Lev_sbd = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1379
          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1380
            (K (mk_Lev_sbd_tac lthy cts Lev_0s Lev_Sucs to_sbd_thmss))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1381
          |> Thm.close_derivation);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1382
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1383
        val Lev_sbd' = mk_specN n Lev_sbd;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1384
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1385
        map (fn i => Lev_sbd' RS mk_conjunctN n i) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1386
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1387
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1388
    val (length_Lev_thms, length_Lev'_thms) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1389
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1390
        fun mk_conjunct i z = HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1391
          HOLogic.mk_eq (mk_size kl, nat));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1392
        val goal = list_all_free (kl :: zs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1393
          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1394
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1395
        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1396
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1397
        val length_Lev = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1398
          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1399
            (K (mk_length_Lev_tac lthy cts Lev_0s Lev_Sucs))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1400
          |> Thm.close_derivation);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1401
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1402
        val length_Lev' = mk_specN (n + 1) length_Lev;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1403
        val length_Levs = map (fn i => length_Lev' RS mk_conjunctN n i RS mp) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1404
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1405
        fun mk_goal i z = fold_rev Logic.all (z :: kl :: nat :: ss) (Logic.mk_implies
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1406
            (HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1407
            HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1408
        val goals = map2 mk_goal ks zs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1409
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1410
        val length_Levs' = map2 (fn goal => fn length_Lev =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1411
          Goal.prove_sorry lthy [] [] goal (K (mk_length_Lev'_tac length_Lev))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1412
          |> Thm.close_derivation) goals length_Levs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1413
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1414
        (length_Levs, length_Levs')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1415
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1416
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1417
    val prefCl_Lev_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1418
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1419
        fun mk_conjunct i z = HOLogic.mk_imp
50058
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49835
diff changeset
  1420
          (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), mk_prefixeq kl_copy kl),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1421
          HOLogic.mk_mem (kl_copy, mk_Lev ss (mk_size kl_copy) i $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1422
        val goal = list_all_free (kl :: kl_copy :: zs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1423
          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1424
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1425
        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1426
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1427
        val prefCl_Lev = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1428
          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1429
            (K (mk_prefCl_Lev_tac lthy cts Lev_0s Lev_Sucs)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1430
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1431
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1432
        val prefCl_Lev' = mk_specN (n + 2) prefCl_Lev;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1433
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1434
        map (fn i => prefCl_Lev' RS mk_conjunctN n i RS mp) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1435
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1436
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1437
    val rv_last_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1438
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1439
        fun mk_conjunct i z i' z_copy = list_exists_free [z_copy]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1440
          (HOLogic.mk_eq
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1441
            (mk_rv ss (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i'])) i $ z,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1442
            mk_InN activeAs z_copy i'));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1443
        val goal = list_all_free (k :: zs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1444
          (Library.foldr1 HOLogic.mk_conj (map2 (fn i => fn z =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1445
            Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1446
              (map2 (mk_conjunct i z) ks zs_copy)) ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1447
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1448
        val cTs = [SOME (certifyT lthy sum_sbdT)];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1449
        val cts = map (SOME o certify lthy) [Term.absfree kl' goal, kl];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1450
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1451
        val rv_last = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1452
          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1453
            (K (mk_rv_last_tac cTs cts rv_Nils rv_Conss)))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1454
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1455
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1456
        val rv_last' = mk_specN (n + 1) rv_last;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1457
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1458
        map (fn i => map (fn i' => rv_last' RS mk_conjunctN n i RS mk_conjunctN n i') ks) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1459
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1460
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1461
    val set_rv_Lev_thmsss = if m = 0 then replicate n (replicate n []) else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1462
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1463
        fun mk_case s sets z z_free = Term.absfree z_free (Library.foldr1 HOLogic.mk_conj
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1464
          (map2 (fn set => fn A => mk_leq (set $ (s $ z)) A) (take m sets) As));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1465
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1466
        fun mk_conjunct i z B = HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1467
          (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_mem (z, B)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1468
          mk_sum_caseN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1469
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1470
        val goal = list_all_free (kl :: zs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1471
          (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct ks zs Bs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1472
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1473
        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1474
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1475
        val set_rv_Lev = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1476
          (Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1477
            (Logic.mk_implies (coalg_prem, HOLogic.mk_Trueprop goal))
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1478
            (K (mk_set_rv_Lev_tac lthy m cts Lev_0s Lev_Sucs rv_Nils rv_Conss
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1479
              coalg_set_thmss from_to_sbd_thmss)))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1480
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1481
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1482
        val set_rv_Lev' = mk_specN (n + 1) set_rv_Lev;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1483
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1484
        map (fn i => map (fn i' =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1485
          split_conj_thm (if n = 1 then set_rv_Lev' RS mk_conjunctN n i RS mp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1486
            else set_rv_Lev' RS mk_conjunctN n i RS mp RSN
52904
traytel
parents: 52839
diff changeset
  1487
              (2, @{thm sum_case_weak_cong} RS iffD1) RS
traytel
parents: 52839
diff changeset
  1488
              (mk_sum_casesN n i' RS iffD1))) ks) ks
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1489
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1490
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1491
    val set_Lev_thmsss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1492
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1493
        fun mk_conjunct i z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1494
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1495
            fun mk_conjunct' i' sets s z' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1496
              let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1497
                fun mk_conjunct'' i'' set z'' = HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1498
                  (HOLogic.mk_mem (z'', set $ (s $ z')),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1499
                    HOLogic.mk_mem (mk_append (kl,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1500
                      HOLogic.mk_list sum_sbdT [mk_InN sbdTs (mk_to_sbd s z' i' i'' $ z'') i'']),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1501
                      mk_Lev ss (HOLogic.mk_Suc nat) i $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1502
              in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1503
                HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z' i'),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1504
                  (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct'' ks (drop m sets) zs_copy2)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1505
              end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1506
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1507
            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1508
              Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct' ks setssAs ss zs_copy))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1509
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1510
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1511
        val goal = list_all_free (kl :: zs @ zs_copy @ zs_copy2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1512
          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1513
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1514
        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1515
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1516
        val set_Lev = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1517
          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1518
            (K (mk_set_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1519
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1520
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1521
        val set_Lev' = mk_specN (3 * n + 1) set_Lev;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1522
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1523
        map (fn i => map (fn i' => map (fn i'' => set_Lev' RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1524
          mk_conjunctN n i RS mp RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1525
          mk_conjunctN n i' RS mp RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1526
          mk_conjunctN n i'' RS mp) ks) ks) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1527
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1528
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1529
    val set_image_Lev_thmsss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1530
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1531
        fun mk_conjunct i z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1532
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1533
            fun mk_conjunct' i' sets =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1534
              let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1535
                fun mk_conjunct'' i'' set s z'' = HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1536
                  (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z'' i''),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1537
                  HOLogic.mk_mem (k, mk_image (mk_to_sbd s z'' i'' i') $ (set $ (s $ z''))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1538
              in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1539
                HOLogic.mk_imp (HOLogic.mk_mem
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1540
                  (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i']),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1541
                    mk_Lev ss (HOLogic.mk_Suc nat) i $ z),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1542
                  (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct'' ks sets ss zs_copy)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1543
              end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1544
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1545
            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1546
              Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct' ks (drop m setssAs')))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1547
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1548
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1549
        val goal = list_all_free (kl :: k :: zs @ zs_copy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1550
          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1551
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1552
        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1553
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1554
        val set_image_Lev = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1555
          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1556
            (K (mk_set_image_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1557
              from_to_sbd_thmss to_sbd_inj_thmss)))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1558
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1559
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1560
        val set_image_Lev' = mk_specN (2 * n + 2) set_image_Lev;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1561
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1562
        map (fn i => map (fn i' => map (fn i'' => set_image_Lev' RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1563
          mk_conjunctN n i RS mp RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1564
          mk_conjunctN n i'' RS mp RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1565
          mk_conjunctN n i' RS mp) ks) ks) ks
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1566
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1567
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1568
    val mor_beh_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1569
      Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1570
        (fold_rev Logic.all (As @ Bs @ ss) (Logic.mk_implies (coalg_prem,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1571
          HOLogic.mk_Trueprop (mk_mor Bs ss carTAs strTAs (map (mk_beh ss) ks)))))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1572
        (mk_mor_beh_tac m mor_def mor_cong_thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1573
          beh_defs carT_defs strT_defs isNode_defs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1574
          to_sbd_inj_thmss from_to_sbd_thmss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbd_thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1575
          length_Lev_thms length_Lev'_thms prefCl_Lev_thms rv_last_thmss
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1576
          set_rv_Lev_thmsss set_Lev_thmsss set_image_Lev_thmsss
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  1577
          set_map'ss coalg_set_thmss map_comp_id_thms map_cong0s map_arg_cong_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1578
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1579
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1580
    val timer = time (timer "Behavioral morphism");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1581
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1582
    fun mk_LSBIS As i = mk_lsbis As (map (mk_carT As) ks) strTAs i;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1583
    fun mk_car_final As i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1584
      mk_quotient (mk_carT As i) (mk_LSBIS As i);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1585
    fun mk_str_final As i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1586
      mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1587
        passive_ids @ map (mk_proj o mk_LSBIS As) ks), nth strTAs (i - 1)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1588
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1589
    val car_finalAs = map (mk_car_final As) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1590
    val str_finalAs = map (mk_str_final As) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1591
    val car_finals = map (mk_car_final passive_UNIVs) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1592
    val str_finals = map (mk_str_final passive_UNIVs) ks;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1593
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1594
    val coalgT_set_thmss = map (map (fn thm => coalgT_thm RS thm)) coalg_set_thmss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1595
    val equiv_LSBIS_thms = map (fn thm => coalgT_thm RS thm) equiv_lsbis_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1596
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1597
    val congruent_str_final_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1598
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1599
        fun mk_goal R final_map strT =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1600
          fold_rev Logic.all As (HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1601
            (mk_congruent R (HOLogic.mk_comp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1602
              (Term.list_comb (final_map, passive_ids @ map (mk_proj o mk_LSBIS As) ks), strT))));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1603
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1604
        val goals = map3 mk_goal (map (mk_LSBIS As) ks) final_maps strTAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1605
      in
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1606
        map4 (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1607
          Goal.prove_sorry lthy [] [] goal
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1608
            (K (mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBIS_thms))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1609
          |> Thm.close_derivation)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1610
        goals lsbisE_thms map_comp_id_thms map_cong0s
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1611
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1612
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1613
    val coalg_final_thm = Goal.prove_sorry lthy [] [] (fold_rev Logic.all As
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1614
      (HOLogic.mk_Trueprop (mk_coalg As car_finalAs str_finalAs)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1615
      (K (mk_coalg_final_tac m coalg_def congruent_str_final_thms equiv_LSBIS_thms
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  1616
        set_map'ss coalgT_set_thmss))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1617
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1618
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1619
    val mor_T_final_thm = Goal.prove_sorry lthy [] [] (fold_rev Logic.all As
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1620
      (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finalAs str_finalAs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1621
        (map (mk_proj o mk_LSBIS As) ks))))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1622
      (K (mk_mor_T_final_tac mor_def congruent_str_final_thms equiv_LSBIS_thms))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1623
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1624
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1625
    val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1626
    val in_car_final_thms = map (fn mor_image' => mor_image' OF
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1627
      [tcoalg_thm RS mor_final_thm, UNIV_I]) mor_image'_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1628
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1629
    val timer = time (timer "Final coalgebra");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1630
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1631
    val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1632
      lthy
49169
937a0fadddfb honor mixfix specifications
blanchet
parents: 49134
diff changeset
  1633
      |> fold_map4 (fn b => fn mx => fn car_final => fn in_car_final =>
49835
31f32ec4d766 discontinued typedef with alternative name;
wenzelm
parents: 49833
diff changeset
  1634
        typedef (b, params, mx) car_final NONE
49169
937a0fadddfb honor mixfix specifications
blanchet
parents: 49134
diff changeset
  1635
          (EVERY' [rtac exI, rtac in_car_final] 1)) bs mixfixes car_finals in_car_final_thms
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1636
      |>> apsnd split_list o split_list;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1637
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1638
    val Ts = map (fn name => Type (name, params')) T_names;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1639
    fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1640
    val Ts' = mk_Ts passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1641
    val Ts'' = mk_Ts passiveCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1642
    val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> treeQT)) T_glob_infos Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1643
    val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, treeQT --> T)) T_glob_infos Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1644
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1645
    val Reps = map #Rep T_loc_infos;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1646
    val Rep_injects = map #Rep_inject T_loc_infos;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1647
    val Abs_inverses = map #Abs_inverse T_loc_infos;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1648
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1649
    val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1650
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1651
    val UNIVs = map HOLogic.mk_UNIV Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1652
    val FTs = mk_FTs (passiveAs @ Ts);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1653
    val FTs' = mk_FTs (passiveBs @ Ts);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1654
    val prodTs = map (HOLogic.mk_prodT o `I) Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1655
    val prodFTs = mk_FTs (passiveAs @ prodTs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1656
    val FTs_setss = mk_setss (passiveAs @ Ts);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1657
    val prodFT_setss = mk_setss (passiveAs @ prodTs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1658
    val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1659
    val map_FT_nths = map2 (fn Ds =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1660
      mk_map_of_bnf Ds (passiveAs @ prodTs) (passiveAs @ Ts)) Dss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1661
    val fstsTs = map fst_const prodTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1662
    val sndsTs = map snd_const prodTs;
52923
traytel
parents: 52913
diff changeset
  1663
    val dtorTs = map2 (curry op -->) Ts FTs;
traytel
parents: 52913
diff changeset
  1664
    val ctorTs = map2 (curry op -->) FTs Ts;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1665
    val unfold_fTs = map2 (curry op -->) activeAs Ts;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1666
    val corec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) sum_sTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1667
    val corec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1668
    val corec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls_rev;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1669
    val corec_Inls = map (Term.subst_atomic_types (activeBs ~~ Ts)) Inls;
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1670
    val corec_UNIVs = map2 (HOLogic.mk_UNIV oo curry mk_sumT) Ts activeAs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1671
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1672
    val ((((((((((((((Jzs, Jzs'), (Jz's, Jz's')), Jzs_copy), Jz's_copy), Jzs1), Jzs2), Jpairs),
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1673
      FJzs), TRs), unfold_fs), unfold_fs_copy), corec_ss), phis), names_lthy) = names_lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1674
      |> mk_Frees' "z" Ts
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1675
      ||>> mk_Frees' "y" Ts'
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1676
      ||>> mk_Frees "z'" Ts
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1677
      ||>> mk_Frees "y'" Ts'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1678
      ||>> mk_Frees "z1" Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1679
      ||>> mk_Frees "z2" Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1680
      ||>> mk_Frees "j" (map2 (curry HOLogic.mk_prodT) Ts Ts')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1681
      ||>> mk_Frees "x" prodFTs
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  1682
      ||>> mk_Frees "r" (map (mk_relT o `I) Ts)
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1683
      ||>> mk_Frees "f" unfold_fTs
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1684
      ||>> mk_Frees "g" unfold_fTs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1685
      ||>> mk_Frees "s" corec_sTs
49592
b859a02c1150 fixed "rels" + split them into injectivity and distinctness
blanchet
parents: 49591
diff changeset
  1686
      ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1687
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1688
    fun dtor_bind i = Binding.suffix_name ("_" ^ dtorN) (nth bs (i - 1));
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1689
    val dtor_name = Binding.name_of o dtor_bind;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1690
    val dtor_def_bind = rpair [] o Thm.def_binding o dtor_bind;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1691
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1692
    fun dtor_spec i rep str map_FT dtorT Jz Jz' =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1693
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1694
        val lhs = Free (dtor_name i, dtorT);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1695
        val rhs = Term.absfree Jz'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1696
          (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1697
            (str $ (rep $ Jz)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1698
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1699
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1700
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1701
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1702
    val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1703
      lthy
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1704
      |> fold_map7 (fn i => fn rep => fn str => fn mapx => fn dtorT => fn Jz => fn Jz' =>
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1705
        Specification.definition (SOME (dtor_bind i, NONE, NoSyn),
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1706
          (dtor_def_bind i, dtor_spec i rep str mapx dtorT Jz Jz')))
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1707
        ks Rep_Ts str_finals map_FTs dtorTs Jzs Jzs'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1708
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1709
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1710
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1711
    val phi = Proof_Context.export_morphism lthy_old lthy;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1712
    fun mk_dtors passive =
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
  1713
      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1714
        Morphism.term phi) dtor_frees;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1715
    val dtors = mk_dtors passiveAs;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1716
    val dtor's = mk_dtors passiveBs;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1717
    val dtor_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) dtor_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1718
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1719
    val coalg_final_set_thmss = map (map (fn thm => coalg_final_thm RS thm)) coalg_set_thmss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1720
    val (mor_Rep_thm, mor_Abs_thm) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1721
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1722
        val mor_Rep =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1723
          Goal.prove_sorry lthy [] []
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1724
            (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts))
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1725
            (mk_mor_Rep_tac m (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1726
              map_comp_id_thms map_cong0L_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1727
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1728
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1729
        val mor_Abs =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1730
          Goal.prove_sorry lthy [] []
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1731
            (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts))
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1732
            (mk_mor_Abs_tac (mor_def :: dtor_defs) Abs_inverses)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1733
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1734
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1735
        (mor_Rep, mor_Abs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1736
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1737
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1738
    val timer = time (timer "dtor definitions & thms");
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1739
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1740
    fun unfold_bind i = Binding.suffix_name ("_" ^ dtor_unfoldN) (nth bs (i - 1));
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1741
    val unfold_name = Binding.name_of o unfold_bind;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1742
    val unfold_def_bind = rpair [] o Thm.def_binding o unfold_bind;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1743
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1744
    fun unfold_spec i T AT abs f z z' =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1745
      let
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1746
        val unfoldT = Library.foldr (op -->) (sTs, AT --> T);
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1747
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1748
        val lhs = Term.list_comb (Free (unfold_name i, unfoldT), ss);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1749
        val rhs = Term.absfree z' (abs $ (f $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1750
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1751
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1752
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1753
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1754
    val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1755
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1756
      |> fold_map7 (fn i => fn T => fn AT => fn abs => fn f => fn z => fn z' =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1757
        Specification.definition
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1758
          (SOME (unfold_bind i, NONE, NoSyn), (unfold_def_bind i, unfold_spec i T AT abs f z z')))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1759
          ks Ts activeAs Abs_Ts (map (fn i => HOLogic.mk_comp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1760
            (mk_proj (mk_LSBIS passive_UNIVs i), mk_beh ss i)) ks) zs zs'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1761
      |>> apsnd split_list o split_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1762
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1763
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1764
    val phi = Proof_Context.export_morphism lthy_old lthy;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1765
    val unfolds = map (Morphism.term phi) unfold_frees;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1766
    val unfold_names = map (fst o dest_Const) unfolds;
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1767
    fun mk_unfolds passives actives =
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1768
      map3 (fn name => fn T => fn active =>
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1769
        Const (name, Library.foldr (op -->)
52923
traytel
parents: 52913
diff changeset
  1770
          (map2 (curry op -->) actives (mk_FTs (passives @ actives)), active --> T)))
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1771
      unfold_names (mk_Ts passives) actives;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1772
    fun mk_unfold Ts ss i = Term.list_comb (Const (nth unfold_names (i - 1), Library.foldr (op -->)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1773
      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1774
    val unfold_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) unfold_def_frees;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1775
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1776
    val mor_unfold_thm =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1777
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1778
        val Abs_inverses' = map2 (curry op RS) in_car_final_thms Abs_inverses;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1779
        val morEs' = map (fn thm =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1780
          (thm OF [tcoalg_thm RS mor_final_thm, UNIV_I]) RS sym) morE_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1781
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1782
        Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1783
          (fold_rev Logic.all ss
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1784
            (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks))))
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1785
          (K (mk_mor_unfold_tac m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs'
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1786
            map_comp_id_thms map_cong0s))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1787
        |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1788
      end;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1789
    val dtor_unfold_thms = map (fn thm => (thm OF [mor_unfold_thm, UNIV_I]) RS sym) morE_thms;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1790
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1791
    val (raw_coind_thms, raw_coind_thm) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1792
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1793
        val prem = HOLogic.mk_Trueprop (mk_sbis passive_UNIVs UNIVs dtors TRs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1794
        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1795
          (map2 (fn R => fn T => mk_leq R (Id_const T)) TRs Ts));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1796
        val goal = fold_rev Logic.all TRs (Logic.mk_implies (prem, concl));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1797
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1798
        `split_conj_thm (Goal.prove_sorry lthy [] [] goal
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1799
          (K (mk_raw_coind_tac bis_def bis_cong_thm bis_O_thm bis_converse_thm bis_Gr_thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1800
            tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1801
            lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1802
          |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1803
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1804
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1805
    val unique_mor_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1806
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1807
        val prems = [HOLogic.mk_Trueprop (mk_coalg passive_UNIVs Bs ss), HOLogic.mk_Trueprop
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1808
          (HOLogic.mk_conj (mk_mor Bs ss UNIVs dtors unfold_fs,
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1809
            mk_mor Bs ss UNIVs dtors unfold_fs_copy))];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1810
        fun mk_fun_eq B f g z = HOLogic.mk_imp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1811
          (HOLogic.mk_mem (z, B), HOLogic.mk_eq (f $ z, g $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1812
        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1813
          (map4 mk_fun_eq Bs unfold_fs unfold_fs_copy zs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1814
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1815
        val unique_mor = Goal.prove_sorry lthy [] []
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1816
          (fold_rev Logic.all (Bs @ ss @ unfold_fs @ unfold_fs_copy @ zs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1817
            (Logic.list_implies (prems, unique)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1818
          (K (mk_unique_mor_tac raw_coind_thms bis_image2_thm))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1819
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1820
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1821
        map (fn thm => conjI RSN (2, thm RS mp)) (split_conj_thm unique_mor)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1822
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1823
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1824
    val (unfold_unique_mor_thms, unfold_unique_mor_thm) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1825
      let
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1826
        val prem = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors unfold_fs);
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1827
        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_unfold Ts ss i);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1828
        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1829
          (map2 mk_fun_eq unfold_fs ks));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1830
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1831
        val bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1832
        val mor_thm = mor_comp_thm OF [tcoalg_thm RS mor_final_thm, mor_Abs_thm];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1833
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1834
        val unique_mor = Goal.prove_sorry lthy [] []
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1835
          (fold_rev Logic.all (ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1836
          (K (mk_unfold_unique_mor_tac raw_coind_thms bis_thm mor_thm unfold_defs))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1837
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1838
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1839
        `split_conj_thm unique_mor
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1840
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1841
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1842
    val (dtor_unfold_unique_thms, dtor_unfold_unique_thm) = `split_conj_thm (split_conj_prems n
52904
traytel
parents: 52839
diff changeset
  1843
      (mor_UNIV_thm RS iffD2 RS unfold_unique_mor_thm));
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1844
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1845
    val unfold_dtor_thms = map (fn thm => mor_id_thm RS thm RS sym) unfold_unique_mor_thms;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1846
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1847
    val unfold_o_dtor_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1848
      let
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1849
        val mor = mor_comp_thm OF [mor_str_thm, mor_unfold_thm];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1850
      in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1851
        map2 (fn unique => fn unfold_ctor =>
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1852
          trans OF [mor RS unique, unfold_ctor]) unfold_unique_mor_thms unfold_dtor_thms
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1853
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1854
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1855
    val timer = time (timer "unfold definitions & thms");
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1856
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1857
    val map_dtors = map2 (fn Ds => fn bnf =>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1858
      Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf,
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1859
        map HOLogic.id_const passiveAs @ dtors)) Dss bnfs;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1860
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1861
    fun ctor_bind i = Binding.suffix_name ("_" ^ ctorN) (nth bs (i - 1));
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1862
    val ctor_name = Binding.name_of o ctor_bind;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1863
    val ctor_def_bind = rpair [] o Thm.def_binding o ctor_bind;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1864
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1865
    fun ctor_spec i ctorT =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1866
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1867
        val lhs = Free (ctor_name i, ctorT);
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1868
        val rhs = mk_unfold Ts map_dtors i;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1869
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1870
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1871
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1872
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1873
    val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
  1874
      lthy
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1875
      |> fold_map2 (fn i => fn ctorT =>
49311
blanchet
parents: 49308
diff changeset
  1876
        Specification.definition
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1877
          (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i ctorT))) ks ctorTs
49311
blanchet
parents: 49308
diff changeset
  1878
      |>> apsnd split_list o split_list
blanchet
parents: 49308
diff changeset
  1879
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
  1880
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1881
    val phi = Proof_Context.export_morphism lthy_old lthy;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1882
    fun mk_ctors params =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1883
      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1884
        ctor_frees;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1885
    val ctors = mk_ctors params';
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1886
    val ctor_defs = map (Morphism.thm phi) ctor_def_frees;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1887
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1888
    val ctor_o_dtor_thms = map2 (fold_thms lthy o single) ctor_defs unfold_o_dtor_thms;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1889
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1890
    val dtor_o_ctor_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1891
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1892
        fun mk_goal dtor ctor FT =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1893
         mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1894
        val goals = map3 mk_goal dtors ctors FTs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1895
      in
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1896
        map5 (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_cong0L =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1897
          Goal.prove_sorry lthy [] [] goal
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1898
            (mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_cong0L unfold_o_dtor_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1899
          |> Thm.close_derivation)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1900
          goals ctor_defs dtor_unfold_thms map_comp_id_thms map_cong0L_thms
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1901
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1902
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1903
    val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1904
    val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1905
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1906
    val bij_dtor_thms =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1907
      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1908
    val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1909
    val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1910
    val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1911
    val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1912
    val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1913
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1914
    val bij_ctor_thms =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1915
      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1916
    val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1917
    val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1918
    val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1919
    val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1920
    val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1921
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1922
    fun mk_ctor_dtor_unfold_like_thm dtor_inject dtor_ctor unfold =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1923
      iffD1 OF [dtor_inject, trans OF [unfold, dtor_ctor RS sym]];
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1924
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1925
    val ctor_dtor_unfold_thms =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1926
      map3 mk_ctor_dtor_unfold_like_thm dtor_inject_thms dtor_ctor_thms dtor_unfold_thms;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1927
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1928
    val timer = time (timer "ctor definitions & thms");
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1929
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1930
    val corec_Inl_sum_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1931
      let
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1932
        val mor = mor_comp_thm OF [mor_sum_case_thm, mor_unfold_thm];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1933
      in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1934
        map2 (fn unique => fn unfold_dtor =>
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1935
          trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1936
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1937
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1938
    fun corec_bind i = Binding.suffix_name ("_" ^ dtor_corecN) (nth bs (i - 1));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1939
    val corec_name = Binding.name_of o corec_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1940
    val corec_def_bind = rpair [] o Thm.def_binding o corec_bind;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1941
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1942
    val corec_strs =
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1943
      map3 (fn dtor => fn sum_s => fn mapx =>
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1944
        mk_sum_case
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1945
          (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ corec_Inls), dtor), sum_s))
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1946
      dtors corec_ss corec_maps;
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1947
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1948
    fun corec_spec i T AT =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1949
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1950
        val corecT = Library.foldr (op -->) (corec_sTs, AT --> T);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1951
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1952
        val lhs = Term.list_comb (Free (corec_name i, corecT), corec_ss);
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1953
        val rhs = HOLogic.mk_comp (mk_unfold Ts corec_strs i, Inr_const T AT);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1954
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1955
        mk_Trueprop_eq (lhs, rhs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1956
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1957
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1958
    val ((corec_frees, (_, corec_def_frees)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
  1959
      lthy
blanchet
parents: 49308
diff changeset
  1960
      |> fold_map3 (fn i => fn T => fn AT =>
blanchet
parents: 49308
diff changeset
  1961
        Specification.definition
blanchet
parents: 49308
diff changeset
  1962
          (SOME (corec_bind i, NONE, NoSyn), (corec_def_bind i, corec_spec i T AT)))
blanchet
parents: 49308
diff changeset
  1963
          ks Ts activeAs
blanchet
parents: 49308
diff changeset
  1964
      |>> apsnd split_list o split_list
blanchet
parents: 49308
diff changeset
  1965
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
  1966
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1967
    val phi = Proof_Context.export_morphism lthy_old lthy;
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
  1968
    val corecs = map (Morphism.term phi) corec_frees;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
  1969
    val corec_names = map (fst o dest_Const) corecs;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
  1970
    fun mk_corec ss i = Term.list_comb (Const (nth corec_names (i - 1), Library.foldr (op -->)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1971
      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1972
    val corec_defs = map (Morphism.thm phi) corec_def_frees;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1973
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1974
    val sum_cases =
49255
2ecc533d6697 use balanced sums for constructors (to gracefully handle 100 constructors or more)
blanchet
parents: 49231
diff changeset
  1975
      map2 (fn T => fn i => mk_sum_case (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1976
    val dtor_corec_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1977
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1978
        fun mk_goal i corec_s corec_map dtor z =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1979
          let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1980
            val lhs = dtor $ (mk_corec corec_ss i $ z);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1981
            val rhs = Term.list_comb (corec_map, passive_ids @ sum_cases) $ (corec_s $ z);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1982
          in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
  1983
            fold_rev Logic.all (z :: corec_ss) (mk_Trueprop_eq (lhs, rhs))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1984
          end;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1985
        val goals = map5 mk_goal ks corec_ss corec_maps_rev dtors zs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1986
      in
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1987
        map3 (fn goal => fn unfold => fn map_cong0 =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1988
          Goal.prove_sorry lthy [] [] goal
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1989
            (mk_corec_tac m corec_defs unfold map_cong0 corec_Inl_sum_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1990
          |> Thm.close_derivation)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1991
        goals dtor_unfold_thms map_cong0s
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1992
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1993
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1994
    val corec_unique_mor_thm =
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1995
      let
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1996
        val id_fs = map2 (fn T => fn f => mk_sum_case (HOLogic.id_const T, f)) Ts unfold_fs;
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1997
        val prem = HOLogic.mk_Trueprop (mk_mor corec_UNIVs corec_strs UNIVs dtors id_fs);
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1998
        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_corec corec_ss i);
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1999
        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2000
          (map2 mk_fun_eq unfold_fs ks));
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2001
      in
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2002
        Goal.prove_sorry lthy [] []
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2003
          (fold_rev Logic.all (corec_ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2004
          (mk_corec_unique_mor_tac corec_defs corec_Inl_sum_thms unfold_unique_mor_thm)
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2005
          |> Thm.close_derivation
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2006
      end;
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2007
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2008
    val (dtor_corec_unique_thms, dtor_corec_unique_thm) =
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2009
      `split_conj_thm (split_conj_prems n
52904
traytel
parents: 52839
diff changeset
  2010
        (mor_UNIV_thm RS iffD2 RS corec_unique_mor_thm)
52634
7c4b56bac189 some new lemmas towards getting rid of in_bd BNF property; tuned
traytel
parents: 52545
diff changeset
  2011
        |> Local_Defs.unfold lthy (@{thms o_sum_case o_id id_o o_assoc sum_case_o_inj(1)} @
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2012
           map_ids @ sym_map_comps) OF replicate n @{thm arg_cong2[of _ _ _ _ sum_case, OF refl]});
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2013
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2014
    val ctor_dtor_corec_thms =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2015
      map3 mk_ctor_dtor_unfold_like_thm dtor_inject_thms dtor_ctor_thms dtor_corec_thms;
49231
43d2df313559 generate "fld_unf_corecs" as well
blanchet
parents: 49228
diff changeset
  2016
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2017
    val timer = time (timer "corec definitions & thms");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2018
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2019
    (* TODO: Get rid of strong versions (since these can easily be derived from the weak ones). *)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2020
    val (dtor_map_coinduct_thm, coinduct_params, dtor_coinduct_thm,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2021
         dtor_map_strong_coinduct_thm, dtor_strong_coinduct_thm) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2022
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2023
        val zs = Jzs1 @ Jzs2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2024
        val frees = phis @ zs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2025
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2026
        fun mk_phi strong_eq phi z1 z2 = if strong_eq
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2027
          then Term.absfree (dest_Free z1) (Term.absfree (dest_Free z2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2028
            (HOLogic.mk_disj (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2))))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2029
          else phi;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2030
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2031
        fun phi_rels strong_eq =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2032
          map3 (fn phi => fn z1 => fn z2 => mk_phi strong_eq phi z1 z2) phis Jzs1 Jzs2;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2033
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2034
        val rels = map (Term.subst_atomic_types ((activeAs ~~ Ts) @ (activeBs ~~ Ts))) relsAsBs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2035
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2036
        fun mk_concl phi z1 z2 = HOLogic.mk_imp (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2037
        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2038
          (map3 mk_concl phis Jzs1 Jzs2));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2039
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2040
        fun mk_rel_prem strong_eq phi dtor rel Jz Jz_copy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2041
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2042
            val concl = Term.list_comb (rel, map HOLogic.eq_const passiveAs @ phi_rels strong_eq) $
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2043
              (dtor $ Jz) $ (dtor $ Jz_copy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2044
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2045
            HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2046
              (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2047
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2048
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2049
        val rel_prems = map5 (mk_rel_prem false) phis dtors rels Jzs Jzs_copy;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2050
        val rel_strong_prems = map5 (mk_rel_prem true) phis dtors rels Jzs Jzs_copy;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2051
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2052
        val dtor_coinduct_goal =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2053
          fold_rev Logic.all frees (Logic.list_implies (rel_prems, concl));
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2054
        val coinduct_params = rev (Term.add_tfrees dtor_coinduct_goal []);
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2055
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2056
        val dtor_coinduct =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2057
          Goal.prove_sorry lthy [] [] dtor_coinduct_goal
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2058
            (K (mk_dtor_coinduct_tac m raw_coind_thm bis_rel_thm rel_congs))
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2059
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2060
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2061
        fun mk_prem strong_eq phi dtor map_nth sets Jz Jz_copy FJz =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2062
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2063
            val xs = [Jz, Jz_copy];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2064
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2065
            fun mk_map_conjunct nths x =
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2066
              HOLogic.mk_eq (Term.list_comb (map_nth, passive_ids @ nths) $ FJz, dtor $ x);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2067
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2068
            fun mk_set_conjunct set phi z1 z2 =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2069
              list_all_free [z1, z2]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2070
                (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (z1, z2), set $ FJz),
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2071
                  mk_phi strong_eq phi z1 z2 $ z1 $ z2));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2072
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2073
            val concl = list_exists_free [FJz] (HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2074
              (Library.foldr1 HOLogic.mk_conj (map2 mk_map_conjunct [fstsTs, sndsTs] xs),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2075
              Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2076
                (map4 mk_set_conjunct (drop m sets) phis Jzs1 Jzs2)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2077
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2078
            fold_rev Logic.all xs (Logic.mk_implies
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2079
              (HOLogic.mk_Trueprop (Term.list_comb (phi, xs)), HOLogic.mk_Trueprop concl))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2080
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2081
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2082
        fun mk_prems strong_eq =
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2083
          map7 (mk_prem strong_eq) phis dtors map_FT_nths prodFT_setss Jzs Jzs_copy FJzs;
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2084
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2085
        val prems = mk_prems false;
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2086
        val strong_prems = mk_prems true;
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2087
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2088
        val dtor_map_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (prems, concl));
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2089
        val dtor_map_coinduct =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2090
          Goal.prove_sorry lthy [] [] dtor_map_coinduct_goal
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2091
            (K (mk_dtor_map_coinduct_tac m ks raw_coind_thm bis_def))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2092
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2093
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2094
        val cTs = map (SOME o certifyT lthy o TFree) coinduct_params;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2095
        val cts = map3 (SOME o certify lthy ooo mk_phi true) phis Jzs1 Jzs2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2096
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2097
        val dtor_strong_coinduct = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2098
          (Goal.prove_sorry lthy [] []
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2099
            (fold_rev Logic.all zs (Logic.list_implies (rel_strong_prems, concl)))
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2100
            (K (mk_dtor_strong_coinduct_tac lthy m cTs cts dtor_coinduct rel_monos rel_eqs)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2101
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2102
49581
4e5bd3883429 renamed "dtor_coinduct" etc. to "dtor_map_coinduct"
blanchet
parents: 49580
diff changeset
  2103
        val dtor_map_strong_coinduct = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2104
          (Goal.prove_sorry lthy [] []
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  2105
            (fold_rev Logic.all zs (Logic.list_implies (strong_prems, concl)))
49581
4e5bd3883429 renamed "dtor_coinduct" etc. to "dtor_map_coinduct"
blanchet
parents: 49580
diff changeset
  2106
            (K (mk_dtor_map_strong_coinduct_tac ks cTs cts dtor_map_coinduct bis_def
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
  2107
              (tcoalg_thm RS bis_Id_on_thm))))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2108
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2109
      in
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2110
        (dtor_map_coinduct, rev (Term.add_tfrees dtor_map_coinduct_goal []),
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2111
         dtor_coinduct, dtor_map_strong_coinduct, dtor_strong_coinduct)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2112
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2113
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2114
    val timer = time (timer "coinduction");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2115
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2116
    fun mk_dtor_map_DEADID_thm dtor_inject map_id =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2117
      trans OF [iffD2 OF [dtor_inject, id_apply], map_id RS sym];
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2118
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2119
    fun mk_dtor_Jrel_DEADID_thm dtor_inject bnf =
51917
f964a9887713 store proper theorems even for fixed points that have no passive live variables
traytel
parents: 51894
diff changeset
  2120
      trans OF [rel_eq_of_bnf bnf RS @{thm predicate2_eqD}, dtor_inject] RS sym;
f964a9887713 store proper theorems even for fixed points that have no passive live variables
traytel
parents: 51894
diff changeset
  2121
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2122
    val JphiTs = map2 mk_pred2T passiveAs passiveBs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2123
    val prodTsTs' = map2 (curry HOLogic.mk_prodT) Ts Ts';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2124
    val fstsTsTs' = map fst_const prodTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2125
    val sndsTsTs' = map snd_const prodTsTs';
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2126
    val activephiTs = map2 mk_pred2T activeAs activeBs;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2127
    val activeJphiTs = map2 mk_pred2T Ts Ts';
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2128
    val (((Jphis, activephis), activeJphis), names_lthy) = names_lthy
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2129
      |> mk_Frees "R" JphiTs
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2130
      ||>> mk_Frees "S" activephiTs
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2131
      ||>> mk_Frees "JR" activeJphiTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2132
    val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2133
    val in_rels = map in_rel_of_bnf bnfs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2134
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2135
    (*register new codatatypes as BNFs*)
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2136
    val (timer, Jbnfs, (folded_dtor_map_o_thms, folded_dtor_map_thms), folded_dtor_set_thmss',
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2137
      dtor_set_induct_thms, dtor_Jrel_thms, lthy) =
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2138
      if m = 0 then
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2139
        (timer, replicate n DEADID_bnf,
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2140
        map_split (`(mk_pointfree lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_id's),
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2141
        replicate n [], [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs, lthy)
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2142
      else let
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2143
        val fTs = map2 (curry op -->) passiveAs passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2144
        val gTs = map2 (curry op -->) passiveBs passiveCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2145
        val f1Ts = map2 (curry op -->) passiveAs passiveYs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2146
        val f2Ts = map2 (curry op -->) passiveBs passiveYs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2147
        val p1Ts = map2 (curry op -->) passiveXs passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2148
        val p2Ts = map2 (curry op -->) passiveXs passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2149
        val pTs = map2 (curry op -->) passiveXs passiveCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2150
        val uTs = map2 (curry op -->) Ts Ts';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2151
        val B1Ts = map HOLogic.mk_setT passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2152
        val B2Ts = map HOLogic.mk_setT passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2153
        val AXTs = map HOLogic.mk_setT passiveXs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2154
        val XTs = mk_Ts passiveXs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2155
        val YTs = mk_Ts passiveYs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2156
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2157
        val ((((((((((((((((((fs, fs'), fs_copy), gs), us),
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2158
          (Jys, Jys')), (Jys_copy, Jys'_copy)), dtor_set_induct_phiss),
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2159
          B1s), B2s), AXs), f1s), f2s), p1s), p2s), ps), (ys, ys')), (ys_copy, ys'_copy)),
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2160
          names_lthy) = names_lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2161
          |> mk_Frees' "f" fTs
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2162
          ||>> mk_Frees "f" fTs
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2163
          ||>> mk_Frees "g" gTs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2164
          ||>> mk_Frees "u" uTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2165
          ||>> mk_Frees' "b" Ts'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2166
          ||>> mk_Frees' "b" Ts'
49592
b859a02c1150 fixed "rels" + split them into injectivity and distinctness
blanchet
parents: 49591
diff changeset
  2167
          ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2168
          ||>> mk_Frees "B1" B1Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2169
          ||>> mk_Frees "B2" B2Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2170
          ||>> mk_Frees "A" AXTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2171
          ||>> mk_Frees "f1" f1Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2172
          ||>> mk_Frees "f2" f2Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2173
          ||>> mk_Frees "p1" p1Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2174
          ||>> mk_Frees "p2" p2Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2175
          ||>> mk_Frees "p" pTs
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2176
          ||>> mk_Frees' "y" passiveAs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2177
          ||>> mk_Frees' "y" passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2178
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2179
        val map_FTFT's = map2 (fn Ds =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2180
          mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2181
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2182
        fun mk_maps ATs BTs Ts mk_T =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2183
          map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ map mk_T Ts)) Dss bnfs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2184
        fun mk_Fmap mk_const fs Ts Fmap = Term.list_comb (Fmap, fs @ map mk_const Ts);
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2185
        fun mk_map mk_const mk_T Ts fs Ts' dtors mk_maps =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2186
          mk_unfold Ts' (map2 (fn dtor => fn Fmap =>
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2187
            HOLogic.mk_comp (mk_Fmap mk_const fs Ts Fmap, dtor)) dtors (mk_maps Ts mk_T));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2188
        val mk_map_id = mk_map HOLogic.id_const I;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2189
        val mk_mapsAB = mk_maps passiveAs passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2190
        val mk_mapsBC = mk_maps passiveBs passiveCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2191
        val mk_mapsAC = mk_maps passiveAs passiveCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2192
        val mk_mapsAY = mk_maps passiveAs passiveYs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2193
        val mk_mapsBY = mk_maps passiveBs passiveYs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2194
        val mk_mapsXA = mk_maps passiveXs passiveAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2195
        val mk_mapsXB = mk_maps passiveXs passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2196
        val mk_mapsXC = mk_maps passiveXs passiveCs;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2197
        val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2198
        val fs_copy_maps = map (mk_map_id Ts fs_copy Ts' dtors mk_mapsAB) ks;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2199
        val gs_maps = map (mk_map_id Ts' gs Ts'' dtor's mk_mapsBC) ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2200
        val fgs_maps =
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2201
          map (mk_map_id Ts (map2 (curry HOLogic.mk_comp) gs fs) Ts'' dtors mk_mapsAC) ks;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2202
        val Xdtors = mk_dtors passiveXs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2203
        val UNIV's = map HOLogic.mk_UNIV Ts';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2204
        val CUNIVs = map HOLogic.mk_UNIV passiveCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2205
        val UNIV''s = map HOLogic.mk_UNIV Ts'';
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2206
        val dtor''s = mk_dtors passiveCs;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2207
        val f1s_maps = map (mk_map_id Ts f1s YTs dtors mk_mapsAY) ks;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2208
        val f2s_maps = map (mk_map_id Ts' f2s YTs dtor's mk_mapsBY) ks;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2209
        val pid_maps = map (mk_map_id XTs ps Ts'' Xdtors mk_mapsXC) ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2210
        val pfst_Fmaps =
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2211
          map (mk_Fmap fst_const p1s prodTsTs') (mk_mapsXA prodTsTs' (fst o HOLogic.dest_prodT));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2212
        val psnd_Fmaps =
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2213
          map (mk_Fmap snd_const p2s prodTsTs') (mk_mapsXB prodTsTs' (snd o HOLogic.dest_prodT));
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2214
        val p1id_Fmaps = map (mk_Fmap HOLogic.id_const p1s prodTsTs') (mk_mapsXA prodTsTs' I);
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2215
        val p2id_Fmaps = map (mk_Fmap HOLogic.id_const p2s prodTsTs') (mk_mapsXB prodTsTs' I);
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2216
        val pid_Fmaps = map (mk_Fmap HOLogic.id_const ps prodTsTs') (mk_mapsXC prodTsTs' I);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2217
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49538
diff changeset
  2218
        val (dtor_map_thms, map_thms) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2219
          let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2220
            fun mk_goal fs_map map dtor dtor' = fold_rev Logic.all fs
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2221
              (mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_map),
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2222
                HOLogic.mk_comp (Term.list_comb (map, fs @ fs_maps), dtor)));
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2223
            val goals = map4 mk_goal fs_maps map_FTFT's dtors dtor's;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2224
            val cTs = map (SOME o certifyT lthy) FTs';
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2225
            val maps =
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2226
              map5 (fn goal => fn cT => fn unfold => fn map_comp' => fn map_cong0 =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2227
                Goal.prove_sorry lthy [] [] goal
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2228
                  (K (mk_map_tac m n cT unfold map_comp' map_cong0))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2229
                |> Thm.close_derivation)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2230
              goals cTs dtor_unfold_thms map_comp's map_cong0s;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2231
          in
52659
58b87aa4dc3b eliminate duplicated theorems (thanks to "Auto solve_direct" in jEdit)
traytel
parents: 52635
diff changeset
  2232
            map_split (fn thm => (thm RS @{thm comp_eq_dest}, thm)) maps
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2233
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2234
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2235
        val map_comp_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2236
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2237
            val goal = fold_rev Logic.all (fs @ gs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2238
              (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2239
                (map3 (fn fmap => fn gmap => fn fgmap =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2240
                   HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2241
                fs_maps gs_maps fgs_maps)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2242
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2243
            split_conj_thm (Goal.prove_sorry lthy [] [] goal
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2244
              (K (mk_map_comp_tac m n map_thms map_comps map_cong0s dtor_unfold_unique_thm))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2245
              |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2246
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2247
49543
53b3c532a082 renamed low-level "map_unique" to have "ctor" or "dtor" in the name
blanchet
parents: 49542
diff changeset
  2248
        val dtor_map_unique_thm =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2249
          let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2250
            fun mk_prem u map dtor dtor' =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2251
              mk_Trueprop_eq (HOLogic.mk_comp (dtor', u),
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2252
                HOLogic.mk_comp (Term.list_comb (map, fs @ us), dtor));
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2253
            val prems = map4 mk_prem us map_FTFT's dtors dtor's;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2254
            val goal =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2255
              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2256
                (map2 (curry HOLogic.mk_eq) us fs_maps));
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2257
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2258
            Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2259
              (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
52911
fe4c2418f069 tuned tactic;
traytel
parents: 52904
diff changeset
  2260
              (mk_dtor_map_unique_tac dtor_unfold_unique_thm sym_map_comps)
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2261
              |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2262
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2263
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2264
        val timer = time (timer "map functions for the new codatatypes");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2265
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52505
diff changeset
  2266
        val bd = mk_cexp sbd sbd;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2267
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2268
        val timer = time (timer "bounds for the new codatatypes");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2269
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2270
        val setss_by_bnf = map (fn i => map2 (mk_hset dtors i) ls passiveAs) ks;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2271
        val setss_by_bnf' = map (fn i => map2 (mk_hset dtor's i) ls passiveBs) ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2272
        val setss_by_range = transpose setss_by_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2273
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2274
        val dtor_set_thmss =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2275
          let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2276
            fun mk_simp_goal relate pas_set act_sets sets dtor z set =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2277
              relate (set $ z, mk_union (pas_set $ (dtor $ z),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2278
                 Library.foldl1 mk_union
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2279
                   (map2 (fn X => mk_UNION (X $ (dtor $ z))) act_sets sets)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2280
            fun mk_goals eq =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2281
              map2 (fn i => fn sets =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2282
                map4 (fn Fsets =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2283
                  mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets)
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2284
                FTs_setss dtors Jzs sets)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2285
              ls setss_by_range;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2286
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2287
            val le_goals = map
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2288
              (fold_rev Logic.all Jzs o HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2289
              (mk_goals (uncurry mk_leq));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2290
            val set_le_thmss = map split_conj_thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2291
              (map4 (fn goal => fn hset_minimal => fn set_hsets => fn set_hset_hsetss =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2292
                Goal.prove_sorry lthy [] [] goal
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2293
                  (K (mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2294
                |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2295
              le_goals hset_minimal_thms set_hset_thmss' set_hset_hset_thmsss');
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2296
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2297
            val simp_goalss = map (map2 (fn z => fn goal =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2298
                Logic.all z (HOLogic.mk_Trueprop goal)) Jzs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2299
              (mk_goals HOLogic.mk_eq);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2300
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2301
            map4 (map4 (fn goal => fn set_le => fn set_incl_hset => fn set_hset_incl_hsets =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2302
              Goal.prove_sorry lthy [] [] goal
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2303
                (K (mk_dtor_set_tac n set_le set_incl_hset set_hset_incl_hsets))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2304
              |> Thm.close_derivation))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2305
            simp_goalss set_le_thmss set_incl_hset_thmss' set_hset_incl_hset_thmsss'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2306
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2307
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2308
        val timer = time (timer "set functions for the new codatatypes");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2309
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2310
        val colss = map2 (fn j => fn T =>
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2311
          map (fn i => mk_hset_rec dtors nat i j T) ks) ls passiveAs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2312
        val colss' = map2 (fn j => fn T =>
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2313
          map (fn i => mk_hset_rec dtor's nat i j T) ks) ls passiveBs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2314
        val Xcolss = map2 (fn j => fn T =>
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2315
          map (fn i => mk_hset_rec Xdtors nat i j T) ks) ls passiveXs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2316
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2317
        val col_natural_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2318
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2319
            fun mk_col_natural f map z col col' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2320
              HOLogic.mk_eq (mk_image f $ (col $ z), col' $ (map $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2321
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2322
            fun mk_goal f cols cols' = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2323
              (map4 (mk_col_natural f) fs_maps Jzs cols cols'));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2324
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2325
            val goals = map3 mk_goal fs colss colss';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2326
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2327
            val ctss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2328
              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2329
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2330
            val thms =
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2331
              map4 (fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2332
                singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2333
                  (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2334
                    (mk_col_natural_tac cts rec_0s rec_Sucs dtor_map_thms set_map'ss))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2335
                |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2336
              goals ctss hset_rec_0ss' hset_rec_Sucss';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2337
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2338
            map (split_conj_thm o mk_specN n) thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2339
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2340
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2341
        val col_bd_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2342
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2343
            fun mk_col_bd z col = mk_ordLeq (mk_card_of (col $ z)) sbd;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2344
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2345
            fun mk_goal cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2346
              (map2 mk_col_bd Jzs cols));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2347
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2348
            val goals = map mk_goal colss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2349
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2350
            val ctss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2351
              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2352
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2353
            val thms =
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2354
              map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2355
                singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2356
                  (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2357
                    (K (mk_col_bd_tac m j cts rec_0s rec_Sucs
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2358
                      sbd_Card_order sbd_Cinfinite set_sbdss)))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2359
                |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2360
              ls goals ctss hset_rec_0ss' hset_rec_Sucss';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2361
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2362
            map (split_conj_thm o mk_specN n) thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2363
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2364
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2365
        val map_cong0_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2366
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2367
            val cTs = map (SOME o certifyT lthy o
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2368
              Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2369
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2370
            fun mk_prem z set f g y y' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2371
              mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2372
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2373
            fun mk_prems sets z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2374
              Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2375
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2376
            fun mk_map_cong0 sets z fmap gmap =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2377
              HOLogic.mk_imp (mk_prems sets z, HOLogic.mk_eq (fmap $ z, gmap $ z));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2378
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2379
            fun mk_coind_body sets (x, T) z fmap gmap y y_copy =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2380
              HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2381
                (HOLogic.mk_mem (z, HOLogic.mk_Collect (x, T, mk_prems sets z)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2382
                  HOLogic.mk_conj (HOLogic.mk_eq (y, fmap $ z),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2383
                    HOLogic.mk_eq (y_copy, gmap $ z)))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2384
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2385
            fun mk_cphi sets (z' as (x, T)) z fmap gmap y' y y'_copy y_copy =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2386
              HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2387
              |> Term.absfree y'_copy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2388
              |> Term.absfree y'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2389
              |> certify lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2390
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2391
            val cphis =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2392
              map9 mk_cphi setss_by_bnf Jzs' Jzs fs_maps fs_copy_maps Jys' Jys Jys'_copy Jys_copy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2393
49581
4e5bd3883429 renamed "dtor_coinduct" etc. to "dtor_map_coinduct"
blanchet
parents: 49580
diff changeset
  2394
            val coinduct = Drule.instantiate' cTs (map SOME cphis) dtor_map_coinduct_thm;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2395
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2396
            val goal =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2397
              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2398
                (map4 mk_map_cong0 setss_by_bnf Jzs fs_maps fs_copy_maps));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2399
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2400
            val thm = singleton (Proof_Context.export names_lthy lthy)
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2401
              (Goal.prove_sorry lthy [] [] goal
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  2402
              (K (mk_mcong_tac lthy m (rtac coinduct) map_comp's dtor_map_thms map_cong0s set_map'ss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2403
              set_hset_thmss set_hset_hset_thmsss)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2404
              |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2405
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2406
            split_conj_thm thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2407
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2408
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2409
        val B1_ins = map2 (mk_in B1s) setss_by_bnf Ts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2410
        val B2_ins = map2 (mk_in B2s) setss_by_bnf' Ts';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2411
        val thePulls = map4 mk_thePull B1_ins B2_ins f1s_maps f2s_maps;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2412
        val thePullTs = passiveXs @ map2 (curry HOLogic.mk_prodT) Ts Ts';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2413
        val thePull_ins = map2 (mk_in (AXs @ thePulls)) (mk_setss thePullTs) (mk_FTs thePullTs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2414
        val pickFs = map5 mk_pickWP thePull_ins pfst_Fmaps psnd_Fmaps
52923
traytel
parents: 52913
diff changeset
  2415
          (map2 (curry op $) dtors Jzs) (map2 (curry op $) dtor's Jz's);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2416
        val pickF_ss = map3 (fn pickF => fn z => fn z' =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2417
          HOLogic.mk_split (Term.absfree z (Term.absfree z' pickF))) pickFs Jzs' Jz's';
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2418
        val picks = map (mk_unfold XTs pickF_ss) ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2419
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2420
        val wpull_prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2421
          (map8 mk_wpull AXs B1s B2s f1s f2s (replicate m NONE) p1s p2s));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2422
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2423
        val map_eq_thms = map2 (fn simp => fn diff => box_equals OF [diff RS iffD2, simp, simp])
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49538
diff changeset
  2424
          dtor_map_thms dtor_inject_thms;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2425
        val map_wpull_thms = map (fn thm => thm OF
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2426
          (replicate m asm_rl @ replicate n @{thm wpull_thePull})) map_wpulls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2427
        val pickWP_assms_tacs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2428
          map3 mk_pickWP_assms_tac set_incl_hset_thmss set_incl_hin_thmss map_eq_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2429
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2430
        val coalg_thePull_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2431
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2432
            val coalg = HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2433
              (mk_coalg CUNIVs thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2434
            val goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2435
              (Logic.mk_implies (wpull_prem, coalg));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2436
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2437
            Goal.prove_sorry lthy [] [] goal (mk_coalg_thePull_tac m coalg_def map_wpull_thms
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2438
              set_map'ss pickWP_assms_tacs)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2439
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2440
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2441
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2442
        val (mor_thePull_fst_thm, mor_thePull_snd_thm, mor_thePull_pick_thm) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2443
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2444
            val mor_fst = HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2445
              (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p1id_Fmaps pickF_ss)
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2446
                UNIVs dtors fstsTsTs');
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2447
            val mor_snd = HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2448
              (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p2id_Fmaps pickF_ss)
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2449
                UNIV's dtor's sndsTsTs');
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2450
            val mor_pick = HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2451
              (mk_mor thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss)
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2452
                UNIV''s dtor''s (map2 (curry HOLogic.mk_comp) pid_maps picks));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2453
49458
blanchet
parents: 49457
diff changeset
  2454
            val fst_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2455
              (Logic.mk_implies (wpull_prem, mor_fst));
49458
blanchet
parents: 49457
diff changeset
  2456
            val snd_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2457
              (Logic.mk_implies (wpull_prem, mor_snd));
49458
blanchet
parents: 49457
diff changeset
  2458
            val pick_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2459
              (Logic.mk_implies (wpull_prem, mor_pick));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2460
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2461
            (Goal.prove_sorry lthy [] [] fst_goal (mk_mor_thePull_fst_tac m mor_def map_wpull_thms
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2462
              map_comp's pickWP_assms_tacs) |> Thm.close_derivation,
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2463
            Goal.prove_sorry lthy [] [] snd_goal (mk_mor_thePull_snd_tac m mor_def map_wpull_thms
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2464
              map_comp's pickWP_assms_tacs) |> Thm.close_derivation,
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2465
            Goal.prove_sorry lthy [] [] pick_goal (mk_mor_thePull_pick_tac mor_def dtor_unfold_thms
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2466
              map_comp's) |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2467
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2468
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2469
        val pick_col_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2470
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2471
            fun mk_conjunct AX Jpair pick thePull col =
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2472
              HOLogic.mk_imp (HOLogic.mk_mem (Jpair, thePull), mk_leq (col $ (pick $ Jpair)) AX);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2473
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2474
            fun mk_concl AX cols =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2475
              list_all_free Jpairs (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2476
                (map4 (mk_conjunct AX) Jpairs picks thePulls cols));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2477
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2478
            val concls = map2 mk_concl AXs Xcolss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2479
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2480
            val ctss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2481
              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2482
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2483
            val goals =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2484
              map (fn concl => Logic.mk_implies (wpull_prem, HOLogic.mk_Trueprop concl)) concls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2485
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2486
            val thms =
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2487
              map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2488
                singleton (Proof_Context.export names_lthy lthy) (Goal.prove_sorry lthy [] [] goal
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2489
                  (mk_pick_col_tac m j cts rec_0s rec_Sucs dtor_unfold_thms set_map'ss
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2490
                    map_wpull_thms pickWP_assms_tacs))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2491
                |> Thm.close_derivation)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2492
              ls goals ctss hset_rec_0ss' hset_rec_Sucss';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2493
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2494
            map (map (fn thm => thm RS mp) o split_conj_thm o mk_specN n) thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2495
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2496
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2497
        val timer = time (timer "helpers for BNF properties");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2498
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2499
        val map_id_tacs =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2500
          map2 (K oo mk_map_id_tac map_thms) dtor_unfold_unique_thms unfold_dtor_thms;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2501
        val map_comp_tacs = map (fn thm => K (rtac (thm RS sym) 1)) map_comp_thms;
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  2502
        val map_cong0_tacs = map (mk_map_cong0_tac m) map_cong0_thms;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2503
        val set_nat_tacss =
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2504
          map2 (map2 (K oo mk_set_map_tac)) hset_defss (transpose col_natural_thmss);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2505
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2506
        val bd_co_tacs = replicate n (K (mk_bd_card_order_tac sbd_card_order));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2507
        val bd_cinf_tacs = replicate n (K (mk_bd_cinfinite_tac sbd_Cinfinite));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2508
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2509
        val set_bd_tacss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2510
          map2 (map2 (K oo mk_set_bd_tac sbd_Cinfinite)) hset_defss (transpose col_bd_thmss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2511
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2512
        val map_wpull_tacs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2513
          map3 (K ooo mk_wpull_tac m coalg_thePull_thm mor_thePull_fst_thm mor_thePull_snd_thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2514
            mor_thePull_pick_thm) unique_mor_thms (transpose pick_col_thmss) hset_defss;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2515
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2516
        val rel_OO_Grp_tacs = replicate n (K (rtac refl 1));
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49452
diff changeset
  2517
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
  2518
        val tacss = map9 zip_axioms map_id_tacs map_comp_tacs map_cong0_tacs set_nat_tacss
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52634
diff changeset
  2519
          bd_co_tacs bd_cinf_tacs set_bd_tacss map_wpull_tacs rel_OO_Grp_tacs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2520
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2521
        val (hset_dtor_incl_thmss, hset_hset_dtor_incl_thmsss, dtor_hset_induct_thms) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2522
          let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2523
            fun tinst_of dtor =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2524
              map (SOME o certify lthy) (dtor :: remove (op =) dtor dtors);
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2525
            fun tinst_of' dtor = case tinst_of dtor of t :: ts => t :: NONE :: ts;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2526
            val Tinst = map (pairself (certifyT lthy))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2527
              (map Logic.varifyT_global (deads @ allAs) ~~ (deads @ passiveAs @ Ts));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2528
            val set_incl_thmss =
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2529
              map2 (fn dtor => map (singleton (Proof_Context.export names_lthy lthy) o
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2530
                Drule.instantiate' [] (tinst_of' dtor) o
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2531
                Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2532
              dtors set_incl_hset_thmss;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2533
49668
blanchet
parents: 49636
diff changeset
  2534
            val tinst = splice (map (SOME o certify lthy) dtors) (replicate n NONE)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2535
            val set_minimal_thms =
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2536
              map (Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2537
                Drule.zero_var_indexes)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2538
              hset_minimal_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2539
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2540
            val set_set_incl_thmsss =
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2541
              map2 (fn dtor => map (map (singleton (Proof_Context.export names_lthy lthy) o
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2542
                Drule.instantiate' [] (NONE :: tinst_of' dtor) o
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2543
                Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2544
              dtors set_hset_incl_hset_thmsss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2545
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2546
            val set_set_incl_thmsss' = transpose (map transpose set_set_incl_thmsss);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2547
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2548
            val incls =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2549
              maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_thmss @
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2550
                @{thms subset_Collect_iff[OF subset_refl]};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2551
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2552
            fun mk_induct_tinst phis jsets y y' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2553
              map4 (fn phi => fn jset => fn Jz => fn Jz' =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2554
                SOME (certify lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y',
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2555
                  HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz))))))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2556
              phis jsets Jzs Jzs';
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2557
            val dtor_set_induct_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2558
              map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2559
                ((set_minimal
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2560
                  |> Drule.instantiate' [] (mk_induct_tinst phis jsets y y')
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2561
                  |> unfold_thms lthy incls) OF
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2562
                  (replicate n ballI @
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2563
                    maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2564
                |> singleton (Proof_Context.export names_lthy lthy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2565
                |> rule_by_tactic lthy (ALLGOALS (TRY o etac asm_rl)))
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2566
              set_minimal_thms set_set_incl_thmsss' setss_by_range ys ys' dtor_set_induct_phiss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2567
          in
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2568
            (set_incl_thmss, set_set_incl_thmsss, dtor_set_induct_thms)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2569
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2570
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2571
        fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit);
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2572
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2573
        val all_unitTs = replicate live HOLogic.unitT;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2574
        val unitTs = replicate n HOLogic.unitT;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2575
        val unit_funs = replicate n (Term.absdummy HOLogic.unitT HOLogic.unit);
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2576
        fun mk_map_args I =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2577
          map (fn i =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2578
            if member (op =) I i then Term.absdummy HOLogic.unitT (nth ys i)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2579
            else mk_undefined (HOLogic.unitT --> nth passiveAs i))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2580
          (0 upto (m - 1));
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2581
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2582
        fun mk_nat_wit Ds bnf (I, wit) () =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2583
          let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2584
            val passiveI = filter (fn i => i < m) I;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2585
            val map_args = mk_map_args passiveI;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2586
          in
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2587
            Term.absdummy HOLogic.unitT (Term.list_comb
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2588
              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ wit)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2589
          end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2590
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2591
        fun mk_dummy_wit Ds bnf I =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2592
          let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2593
            val map_args = mk_map_args I;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2594
          in
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2595
            Term.absdummy HOLogic.unitT (Term.list_comb
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2596
              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2597
              mk_undefined (mk_T_of_bnf Ds all_unitTs bnf))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2598
          end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2599
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2600
        val nat_witss =
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2601
          map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds)
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2602
            (replicate (nwits_of_bnf bnf) (replicate live HOLogic.unitT)) bnf
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2603
            |> map (fn (I, wit) =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2604
              (I, Lazy.lazy (mk_nat_wit Ds bnf (I, Term.list_comb (wit, map (K HOLogic.unit) I))))))
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2605
          Dss bnfs;
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2606
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2607
        val nat_wit_thmss = map2 (curry op ~~) nat_witss (map wit_thmss_of_bnf bnfs)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2608
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2609
        val Iss = map (map fst) nat_witss;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2610
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2611
        fun filter_wits (I, wit) =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2612
          let val J = filter (fn i => i < m) I;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2613
          in (J, (length J < length I, wit)) end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2614
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2615
        val wit_treess = map_index (fn (i, Is) =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2616
          map_index (finish Iss m [i+m] (i+m)) Is) Iss
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2617
          |> map (minimize_wits o map filter_wits o minimize_wits o flat);
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2618
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2619
        val coind_wit_argsss =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2620
          map (map (tree_to_coind_wits nat_wit_thmss o snd o snd) o filter (fst o snd)) wit_treess;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2621
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2622
        val nonredundant_coind_wit_argsss =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2623
          fold (fn i => fn argsss =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2624
            nth_map (i - 1) (filter_out (fn xs =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2625
              exists (fn ys =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2626
                let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2627
                  val xs' = (map (fst o fst) xs, snd (fst (hd xs)));
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2628
                  val ys' = (map (fst o fst) ys, snd (fst (hd ys)));
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2629
                in
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2630
                  eq_pair (subset (op =)) (eq_set (op =)) (xs', ys') andalso not (fst xs' = fst ys')
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2631
                end)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2632
              (flat argsss)))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2633
            argsss)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2634
          ks coind_wit_argsss;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2635
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2636
        fun prepare_args args =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2637
          let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2638
            val I = snd (fst (hd args));
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2639
            val (dummys, args') =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2640
              map_split (fn i =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2641
                (case find_first (fn arg => fst (fst arg) = i - 1) args of
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2642
                  SOME (_, ((_, wit), thms)) => (NONE, (Lazy.force wit, thms))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2643
                | NONE =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2644
                  (SOME (i - 1), (mk_dummy_wit (nth Dss (i - 1)) (nth bnfs (i - 1)) I, []))))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2645
              ks;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2646
          in
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2647
            ((I, dummys), apsnd flat (split_list args'))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2648
          end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2649
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2650
        fun mk_coind_wits ((I, dummys), (args, thms)) =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2651
          ((I, dummys), (map (fn i => mk_unfold Ts args i $ HOLogic.unit) ks, thms));
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2652
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2653
        val coind_witss =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2654
          maps (map (mk_coind_wits o prepare_args)) nonredundant_coind_wit_argsss;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2655
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2656
        fun mk_coind_wit_thms ((I, dummys), (wits, wit_thms)) =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2657
          let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2658
            fun mk_goal sets y y_copy y'_copy j =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2659
              let
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2660
                fun mk_conjunct set z dummy wit =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2661
                  mk_Ball (set $ z) (Term.absfree y'_copy
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2662
                    (if dummy = NONE orelse member (op =) I (j - 1) then
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2663
                      HOLogic.mk_imp (HOLogic.mk_eq (z, wit),
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2664
                        if member (op =) I (j - 1) then HOLogic.mk_eq (y_copy, y)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2665
                        else @{term False})
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2666
                    else @{term True}));
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2667
              in
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2668
                fold_rev Logic.all (map (nth ys) I @ Jzs) (HOLogic.mk_Trueprop
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2669
                  (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct sets Jzs dummys wits)))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2670
              end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2671
            val goals = map5 mk_goal setss_by_range ys ys_copy ys'_copy ls;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2672
          in
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2673
            map2 (fn goal => fn induct =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2674
              Goal.prove_sorry lthy [] [] goal
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2675
                (mk_coind_wit_tac induct dtor_unfold_thms (flat set_map'ss) wit_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2676
              |> Thm.close_derivation)
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2677
            goals dtor_hset_induct_thms
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2678
            |> map split_conj_thm
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2679
            |> transpose
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2680
            |> map (map_filter (try (fn thm => thm RS bspec RS mp)))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2681
            |> curry op ~~ (map_index Library.I (map (close_wit I) wits))
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2682
            |> filter (fn (_, thms) => length thms = m)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2683
          end;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2684
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2685
        val coind_wit_thms = maps mk_coind_wit_thms coind_witss;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2686
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2687
        val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2688
          (replicate (nwits_of_bnf bnf) Ds)
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2689
          (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2690
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2691
        val ctor_witss =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2692
          map (map (uncurry close_wit o tree_to_ctor_wit ys ctors witss o snd o snd) o
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2693
            filter_out (fst o snd)) wit_treess;
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2694
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2695
        val all_witss =
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2696
          fold (fn ((i, wit), thms) => fn witss =>
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2697
            nth_map i (fn (thms', wits) => (thms @ thms', wit :: wits)) witss)
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2698
          coind_wit_thms (map (pair []) ctor_witss)
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2699
          |> map (apsnd (map snd o minimize_wits));
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2700
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2701
        val wit_tac = mk_wit_tac n dtor_ctor_thms (flat dtor_set_thmss) (maps wit_thms_of_bnf bnfs);
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2702
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2703
        val (Jbnfs, lthy) =
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  2704
          fold_map9 (fn tacs => fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets =>
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  2705
              fn T => fn (thms, wits) => fn lthy =>
51758
55963309557b honor user-specified name for map function
blanchet
parents: 51757
diff changeset
  2706
            bnf_def Dont_Inline (user_policy Note_Some) I tacs (wit_tac thms) (SOME deads) map_b
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  2707
              rel_b set_bs
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  2708
              (((((b, fold_rev Term.absfree fs' mapx), sets), absdummy T bd), wits), NONE) lthy
49434
433dc7e028c8 separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
traytel
parents: 49362
diff changeset
  2709
            |> register_bnf (Local_Theory.full_name lthy b))
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  2710
          tacss bs map_bs rel_bs set_bss fs_maps setss_by_bnf Ts all_witss lthy;
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2711
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2712
        val fold_maps = fold_thms lthy (map (fn bnf =>
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2713
          mk_unabs_def m (map_def_of_bnf bnf RS meta_eq_to_obj_eq)) Jbnfs);
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2714
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2715
        val fold_sets = fold_thms lthy (maps (fn bnf =>
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2716
         map (fn thm => thm RS meta_eq_to_obj_eq) (set_defs_of_bnf bnf)) Jbnfs);
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2717
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2718
        val timer = time (timer "registered new codatatypes as BNFs");
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2719
49544
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
  2720
        val dtor_set_incl_thmss = map (map fold_sets) hset_dtor_incl_thmss;
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
  2721
        val dtor_set_set_incl_thmsss = map (map (map fold_sets)) hset_hset_dtor_incl_thmsss;
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2722
        val dtor_set_induct_thms = map fold_sets dtor_hset_induct_thms;
49104
6defdacd595a generate coinductive witnesses for codatatypes
traytel
parents: 49074
diff changeset
  2723
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
  2724
        val Jrels = map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2725
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2726
        val Jrelphis = map (fn Jrel => Term.list_comb (Jrel, Jphis)) Jrels;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2727
        val relphis = map (fn rel => Term.list_comb (rel, Jphis @ Jrelphis)) rels;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2728
        val in_Jrels = map in_rel_of_bnf Jbnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2729
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49538
diff changeset
  2730
        val folded_dtor_map_thms = map fold_maps dtor_map_thms;
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2731
        val folded_dtor_map_o_thms = map fold_maps map_thms;
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2732
        val folded_dtor_set_thmss = map (map fold_sets) dtor_set_thmss;
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2733
        val folded_dtor_set_thmss' = transpose folded_dtor_set_thmss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2734
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2735
        val dtor_Jrel_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2736
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2737
            fun mk_goal Jz Jz' dtor dtor' Jrelphi relphi = fold_rev Logic.all (Jz :: Jz' :: Jphis)
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2738
              (mk_Trueprop_eq (Jrelphi $ Jz $ Jz', relphi $ (dtor $ Jz) $ (dtor' $ Jz')));
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2739
            val goals = map6 mk_goal Jzs Jz's dtors dtor's Jrelphis relphis;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2740
          in
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2741
            map12 (fn i => fn goal => fn in_rel => fn map_comp => fn map_cong0 =>
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2742
              fn dtor_map => fn dtor_sets => fn dtor_inject => fn dtor_ctor =>
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2743
              fn set_maps => fn dtor_set_incls => fn dtor_set_set_inclss =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2744
              Goal.prove_sorry lthy [] [] goal
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2745
                (K (mk_dtor_rel_tac lthy in_Jrels i in_rel map_comp map_cong0 dtor_map dtor_sets
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2746
                  dtor_inject dtor_ctor set_maps dtor_set_incls dtor_set_set_inclss))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2747
              |> Thm.close_derivation)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2748
            ks goals in_rels map_comp's map_cong0s folded_dtor_map_thms folded_dtor_set_thmss'
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
  2749
              dtor_inject_thms dtor_ctor_thms set_map'ss dtor_set_incl_thmss
49544
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
  2750
              dtor_set_set_incl_thmsss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2751
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2752
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2753
        val timer = time (timer "additional properties");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2754
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2755
        val ls' = if m = 1 then [0] else ls;
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2756
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2757
        val Jbnf_common_notes =
49543
53b3c532a082 renamed low-level "map_unique" to have "ctor" or "dtor" in the name
blanchet
parents: 49542
diff changeset
  2758
          [(dtor_map_uniqueN, [fold_maps dtor_map_unique_thm])] @
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
  2759
          map2 (fn i => fn thm => (mk_dtor_set_inductN i, [thm])) ls' dtor_set_induct_thms
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2760
          |> map (fn (thmN, thms) =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2761
            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2762
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2763
        val Jbnf_notes =
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49538
diff changeset
  2764
          [(dtor_mapN, map single folded_dtor_map_thms),
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49538
diff changeset
  2765
          (dtor_relN, map single dtor_Jrel_thms),
49544
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
  2766
          (dtor_set_inclN, dtor_set_incl_thmss),
49580
040cfb087b3a leave out some internal theorems unless "bnf_note_all" is set
blanchet
parents: 49545
diff changeset
  2767
          (dtor_set_set_inclN, map flat dtor_set_set_incl_thmsss)] @
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  2768
          map2 (fn i => fn thms => (mk_dtor_setN i, map single thms)) ls' folded_dtor_set_thmss
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2769
          |> maps (fn (thmN, thmss) =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2770
            map2 (fn b => fn thms =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2771
              ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2772
            bs thmss)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2773
      in
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2774
       (timer, Jbnfs, (folded_dtor_map_o_thms, folded_dtor_map_thms), folded_dtor_set_thmss',
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2775
         dtor_set_induct_thms, dtor_Jrel_thms,
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2776
         lthy |> Local_Theory.notes (Jbnf_common_notes @ Jbnf_notes) |> snd)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2777
      end;
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2778
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2779
      val dtor_unfold_o_map_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2780
        dtor_unfold_unique_thm folded_dtor_map_o_thms (map (mk_pointfree lthy) dtor_unfold_thms)
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2781
        sym_map_comps map_cong0s;
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2782
      val dtor_corec_o_map_thms = mk_xtor_un_fold_o_map_thms Greatest_FP true m
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2783
        dtor_corec_unique_thm folded_dtor_map_o_thms (map (mk_pointfree lthy) dtor_corec_thms)
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2784
        sym_map_comps map_cong0s;
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2785
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2786
      val passiveABs = map2 (curry HOLogic.mk_prodT) passiveAs passiveBs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2787
      val zip_ranTs = passiveABs @ prodTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2788
      val allJphis = Jphis @ activeJphis;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2789
      val zipFTs = mk_FTs zip_ranTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2790
      val zipTs = map3 (fn T => fn T' => fn FT => T --> T' --> FT) Ts Ts' zipFTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2791
      val zip_zTs = mk_Ts passiveABs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2792
      val (((zips, (abs, abs')), zip_zs), names_lthy) = names_lthy
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2793
        |> mk_Frees "zip" zipTs
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2794
        ||>> mk_Frees' "ab" passiveABs
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2795
        ||>> mk_Frees "z" zip_zTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2796
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2797
      val Iphi_sets =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2798
        map2 (fn phi => fn T => HOLogic.Collect_const T $ HOLogic.mk_split phi) allJphis zip_ranTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2799
      val in_phis = map2 (mk_in Iphi_sets) (mk_setss zip_ranTs) zipFTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2800
      val fstABs = map fst_const passiveABs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2801
      val all_fsts = fstABs @ fstsTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2802
      val map_all_fsts = map2 (fn Ds => fn bnf =>
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2803
        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveAs @ Ts) bnf, all_fsts)) Dss bnfs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2804
      val Jmap_fsts = map2 (fn bnf => fn T => if m = 0 then HOLogic.id_const T
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2805
        else Term.list_comb (mk_map_of_bnf deads passiveABs passiveAs bnf, fstABs)) Jbnfs Ts;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2806
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2807
      val sndABs = map snd_const passiveABs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2808
      val all_snds = sndABs @ sndsTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2809
      val map_all_snds = map2 (fn Ds => fn bnf =>
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2810
        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveBs @ Ts') bnf, all_snds)) Dss bnfs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2811
      val Jmap_snds = map2 (fn bnf => fn T => if m = 0 then HOLogic.id_const T
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2812
        else Term.list_comb (mk_map_of_bnf deads passiveABs passiveBs bnf, sndABs)) Jbnfs Ts;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2813
      val zip_unfolds = map (mk_unfold zip_zTs (map HOLogic.mk_split zips)) ks;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2814
      val zip_setss = map (mk_sets_of_bnf (replicate m deads) (replicate m passiveABs)) Jbnfs
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2815
        |> transpose;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2816
      val in_Jrels = map in_rel_of_bnf Jbnfs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2817
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2818
      val Jrel_coinduct_tac =
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2819
        let
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2820
          fun mk_helper_prem phi in_phi zip x y map map' dtor dtor' =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2821
            let
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2822
              val zipxy = zip $ x $ y;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2823
            in
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2824
              HOLogic.mk_Trueprop (list_all_free [x, y]
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2825
                (HOLogic.mk_imp (phi $ x $ y, HOLogic.mk_conj (HOLogic.mk_mem (zipxy, in_phi),
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2826
                  HOLogic.mk_conj (HOLogic.mk_eq (map $ zipxy, dtor $ x),
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2827
                    HOLogic.mk_eq (map' $ zipxy, dtor' $ y))))))
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2828
            end;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2829
          val helper_prems = map9 mk_helper_prem
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2830
            activeJphis in_phis zips Jzs Jz's map_all_fsts map_all_snds dtors dtor's;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2831
          fun mk_helper_coind_concl fst phi x alt y map zip_unfold =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2832
            HOLogic.mk_imp (list_exists_free [if fst then y else x] (HOLogic.mk_conj (phi $ x $ y,
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2833
              HOLogic.mk_eq (alt, map $ (zip_unfold $ HOLogic.mk_prod (x, y))))),
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2834
            HOLogic.mk_eq (alt, if fst then x else y));
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2835
          val helper_coind1_concl =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2836
            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2837
              (map6 (mk_helper_coind_concl true)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2838
              activeJphis Jzs Jzs_copy Jz's Jmap_fsts zip_unfolds));
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2839
          val helper_coind2_concl =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2840
            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2841
              (map6 (mk_helper_coind_concl false)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2842
              activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds));
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2843
          val helper_coind_tac = mk_rel_coinduct_coind_tac m dtor_map_coinduct_thm ks map_comp's
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2844
            map_cong0s map_arg_cong_thms set_map'ss dtor_unfold_thms folded_dtor_map_thms;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2845
          fun mk_helper_coind_thms vars concl =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2846
            Goal.prove_sorry lthy [] []
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2847
              (fold_rev Logic.all (Jphis @ activeJphis @ vars @ zips)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2848
                (Logic.list_implies (helper_prems, concl)))
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2849
              helper_coind_tac
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2850
            |> Thm.close_derivation
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2851
            |> split_conj_thm;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2852
          val helper_coind1_thms = mk_helper_coind_thms (Jzs @ Jzs_copy) helper_coind1_concl;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2853
          val helper_coind2_thms = mk_helper_coind_thms (Jz's @ Jz's_copy) helper_coind2_concl;
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2854
  
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2855
          fun mk_helper_ind_concl phi ab' ab fst snd z active_phi x y zip_unfold set =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2856
            mk_Ball (set $ z) (Term.absfree ab' (list_all_free [x, y] (HOLogic.mk_imp
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2857
              (HOLogic.mk_conj (active_phi $ x $ y,
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2858
                 HOLogic.mk_eq (z, zip_unfold $ HOLogic.mk_prod (x, y))),
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2859
              phi $ (fst $ ab) $ (snd $ ab)))));
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2860
  
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2861
          val mk_helper_ind_concls =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2862
            map6 (fn Jphi => fn ab' => fn ab => fn fst => fn snd => fn zip_sets =>
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2863
              map6 (mk_helper_ind_concl Jphi ab' ab fst snd)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2864
              zip_zs activeJphis Jzs Jz's zip_unfolds zip_sets)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2865
            Jphis abs' abs fstABs sndABs zip_setss
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2866
            |> map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj);
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2867
  
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2868
          val helper_ind_thmss = if m = 0 then replicate n [] else
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2869
            map3 (fn concl => fn j => fn set_induct =>
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2870
              Goal.prove_sorry lthy [] []
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2871
                (fold_rev Logic.all (Jphis @ activeJphis @ zip_zs @ zips)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2872
                  (Logic.list_implies (helper_prems, concl)))
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2873
                (mk_rel_coinduct_ind_tac m ks dtor_unfold_thms set_map'ss j set_induct)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2874
              |> Thm.close_derivation
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2875
              |> split_conj_thm)
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2876
            mk_helper_ind_concls ls dtor_set_induct_thms
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2877
            |> transpose;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2878
        in
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2879
          mk_rel_coinduct_tac in_rels in_Jrels
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2880
            helper_ind_thmss helper_coind1_thms helper_coind2_thms
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2881
        end;
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2882
      
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2883
      val Jrels = if m = 0 then map HOLogic.eq_const Ts
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2884
        else map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs;
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2885
      val Jrel_coinduct_thm =
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2886
        mk_rel_co_induct_thm Greatest_FP rels activeJphis Jrels Jphis Jzs Jz's dtors dtor's
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2887
          Jrel_coinduct_tac lthy;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2888
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2889
      val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2890
      val dtor_unfold_transfer_thms =
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2891
        mk_un_fold_transfer_thms Greatest_FP rels activephis Jrels Jphis
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2892
          (mk_unfolds passiveAs activeAs) (mk_unfolds passiveBs activeBs)
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2893
          (mk_unfold_transfer_tac m Jrel_coinduct_thm (map map_transfer_of_bnf bnfs)
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2894
            dtor_unfold_thms)
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2895
          lthy;
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2896
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2897
      val timer = time (timer "relator coinduction");
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2898
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2899
      val common_notes =
49582
557302525778 renamed "dtor_rel_coinduct" etc. to "dtor_coinduct"
blanchet
parents: 49581
diff changeset
  2900
        [(dtor_coinductN, [dtor_coinduct_thm]),
49588
9b72d207617b export "dtor_map_coinduct" theorems, since they're used in one example
blanchet
parents: 49585
diff changeset
  2901
        (dtor_map_coinductN, [dtor_map_coinduct_thm]),
9b72d207617b export "dtor_map_coinduct" theorems, since they're used in one example
blanchet
parents: 49585
diff changeset
  2902
        (dtor_map_strong_coinductN, [dtor_map_strong_coinduct_thm]),
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2903
        (dtor_strong_coinductN, [dtor_strong_coinduct_thm]),
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2904
        (rel_coinductN, [Jrel_coinduct_thm]),
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2905
        (dtor_unfold_transferN, dtor_unfold_transfer_thms)]
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2906
        |> map (fn (thmN, thms) =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2907
          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2908
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2909
      val notes =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2910
        [(ctor_dtorN, ctor_dtor_thms),
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49592
diff changeset
  2911
        (ctor_dtor_corecN, ctor_dtor_corec_thms),
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49592
diff changeset
  2912
        (ctor_dtor_unfoldN, ctor_dtor_unfold_thms),
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2913
        (ctor_exhaustN, ctor_exhaust_thms),
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2914
        (ctor_injectN, ctor_inject_thms),
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49592
diff changeset
  2915
        (dtor_corecN, dtor_corec_thms),
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2916
        (dtor_ctorN, dtor_ctor_thms),
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  2917
        (dtor_exhaustN, dtor_exhaust_thms),
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  2918
        (dtor_injectN, dtor_inject_thms),
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49592
diff changeset
  2919
        (dtor_unfoldN, dtor_unfold_thms),
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  2920
        (dtor_unfold_uniqueN, dtor_unfold_unique_thms),
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2921
        (dtor_corec_uniqueN, dtor_corec_unique_thms),
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2922
        (dtor_unfold_o_mapN, dtor_unfold_o_map_thms),
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  2923
        (dtor_corec_o_mapN, dtor_corec_o_map_thms)]
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2924
        |> map (apsnd (map single))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2925
        |> maps (fn (thmN, thmss) =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2926
          map2 (fn b => fn thms =>
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2927
            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2928
          bs thmss)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2929
  in
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2930
    timer;
52328
2f286a2b7f98 [mq]: tuning
blanchet
parents: 52314
diff changeset
  2931
    ({Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors,
52344
blanchet
parents: 52328
diff changeset
  2932
      xtor_co_iterss = transpose [unfolds, corecs],
blanchet
parents: 52328
diff changeset
  2933
      xtor_co_inducts = [dtor_coinduct_thm, dtor_strong_coinduct_thm], dtor_ctors = dtor_ctor_thms,
52312
f461dca57c66 tuned record field names to avoid confusion between low-level and high-level constants/theorems
blanchet
parents: 52298
diff changeset
  2934
      ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms,
52314
9606cf677021 continuation of f461dca57c66
blanchet
parents: 52312
diff changeset
  2935
      xtor_map_thms = folded_dtor_map_thms, xtor_set_thmss = folded_dtor_set_thmss',
52328
2f286a2b7f98 [mq]: tuning
blanchet
parents: 52314
diff changeset
  2936
      xtor_rel_thms = dtor_Jrel_thms,
52839
2c0e1a84dcc7 store relator induction in fp_result
traytel
parents: 52731
diff changeset
  2937
      xtor_co_iter_thmss = transpose [ctor_dtor_unfold_thms, ctor_dtor_corec_thms],
2c0e1a84dcc7 store relator induction in fp_result
traytel
parents: 52731
diff changeset
  2938
      rel_co_induct_thm = Jrel_coinduct_thm},
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49185
diff changeset
  2939
     lthy |> Local_Theory.notes (common_notes @ notes) |> snd)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2940
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2941
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2942
val _ =
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51798
diff changeset
  2943
  Outer_Syntax.local_theory @{command_spec "codatatype"} "define BNF-based coinductive datatypes"
52207
21026c312cc3 tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents: 52090
diff changeset
  2944
    (parse_co_datatype_cmd Greatest_FP construct_gfp);
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49277
diff changeset
  2945
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2946
end;