src/HOL/Tools/BNF/bnf_gfp.ML
author traytel
Thu, 16 Jul 2015 12:23:22 +0200
changeset 60728 26ffdb966759
parent 59936 b8ffc3dc9e24
child 60784 4f590c08fd5d
permissions -rw-r--r--
{r,e,d,f}tac with proper context in BNF
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55061
a0adf838e2d1 adjusted comments
blanchet
parents: 55060
diff changeset
     1
(*  Title:      HOL/Tools/BNF/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 ->
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55770
diff changeset
    14
    BNF_Comp.absT_info list -> 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*)
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55770
diff changeset
    58
fun construct_gfp mixfixes map_bs rel_bs set_bss0 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
53143
ba80154a1118 configuration option to control timing output for (co)datatypes
traytel
parents: 53138
diff changeset
    60
    val time = time lthy;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
    val timer = time (Timer.startRealTimer ());
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
    val 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
    64
    val n = length bnfs; (*active*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
    val ks = 1 upto n;
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
    66
    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
    67
    val ls = 1 upto m;
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
    68
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
    69
    val note_all = Config.get lthy bnf_note_all;
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
    70
    val b_names = map Binding.name_of bs;
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
    71
    val b_name = mk_common_name b_names;
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
    72
    val b = Binding.name b_name;
58241
ff8059e3e803 generate better internal names, with name of the target type in it
blanchet
parents: 58208
diff changeset
    73
ff8059e3e803 generate better internal names, with name of the target type in it
blanchet
parents: 58208
diff changeset
    74
    fun mk_internal_of_b name =
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59819
diff changeset
    75
      Binding.prefix_name (name ^ "_") #> Binding.prefix true b_name #> Binding.concealed;
58241
ff8059e3e803 generate better internal names, with name of the target type in it
blanchet
parents: 58208
diff changeset
    76
    fun mk_internal_b name = mk_internal_of_b name b;
ff8059e3e803 generate better internal names, with name of the target type in it
blanchet
parents: 58208
diff changeset
    77
    fun mk_internal_bs name = map (mk_internal_of_b name) bs;
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
    78
    val external_bs = map2 (Binding.prefix false) b_names bs
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59819
diff changeset
    79
      |> not note_all ? map Binding.concealed;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
    81
    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
    82
    val names_lthy = fold Variable.declare_typ deads lthy;
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
    83
    val passives = map fst (subtract (op = o apsnd TFree) deads resBs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
    (* tvars *)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
    86
    val ((((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs), idxT) = names_lthy
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
    87
      |> variant_tfrees passives
52938
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    88
      ||>> mk_TFrees n
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
    89
      ||>> variant_tfrees passives
52938
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    90
      ||>> mk_TFrees n
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    91
      ||>> mk_TFrees m
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    92
      ||>> mk_TFrees n
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
      ||> fst o mk_TFrees 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
      ||> the_single;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
52938
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    96
    val allAs = passiveAs @ activeAs;
3b3e52d5e66b tuned name generation code (to make it easier to adapt later)
blanchet
parents: 52923
diff changeset
    97
    val allBs' = passiveBs @ activeBs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
    val Ass = replicate n allAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
    val allBs = passiveAs @ activeBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
    val Bss = replicate n allBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
    val allCs = passiveAs @ activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
    val allCs' = passiveBs @ activeCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
    val Css' = replicate n allCs';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
51866
142a82883831 removed dead code
blanchet
parents: 51865
diff changeset
   105
    (* types *)
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   106
    val dead_poss =
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
   107
      map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   108
    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
   109
      | 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
   110
    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
   111
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
    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
   113
    val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
    val FTsAs = mk_FTs allAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
    val FTsBs = mk_FTs allBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
    val FTsCs = mk_FTs allCs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
    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
   118
    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
   119
    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
   120
    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
   121
    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
   122
    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
   123
    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
   124
    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
   125
    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
   126
    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
   127
    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
   128
    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
   129
    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
   130
    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
   131
    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
   132
    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
   133
    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
   134
    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
   135
    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
   136
    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
   137
    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
   138
    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
   139
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
    (* terms *)
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   141
    val mapsAsAs = @{map 4} mk_map_of_bnf Dss Ass Ass bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   142
    val mapsAsBs = @{map 4} mk_map_of_bnf Dss Ass Bss bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   143
    val mapsBsCs' = @{map 4} mk_map_of_bnf Dss Bss Css' bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   144
    val mapsAsCs' = @{map 4} mk_map_of_bnf Dss Ass Css' bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   145
    val map_Inls = @{map 4} mk_map_of_bnf Dss Bss (replicate n (passiveAs @ sumBsAs)) bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   146
    val map_Inls_rev = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ sumBsAs)) Bss bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   147
    val map_fsts = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   148
    val map_snds = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   149
    fun mk_setss Ts = @{map 3} mk_sets_of_bnf (map (replicate live) Dss)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
      (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
   151
    val setssAs = mk_setss allAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
    val setssAs' = transpose setssAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   153
    val bis_setss = mk_setss (passiveAs @ RTs);
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   154
    val relsAsBs = @{map 4} mk_rel_of_bnf Dss Ass Bss bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   155
    val bds = @{map 3} mk_bd_of_bnf Dss Ass bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
    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
   157
    val sum_bdT = fst (dest_relT (fastype_of sum_bd));
56516
a13c2ccc160b more accurate type arguments for intermeadiate typedefs
traytel
parents: 56348
diff changeset
   158
    val (sum_bdT_params, sum_bdT_params') = `(map TFree) (Term.add_tfreesT sum_bdT []);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
   160
    val ((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2), z's), (ys, ys')),
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   161
      Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
   162
      self_fs), gs), all_gs), xFs), yFs), yFs_copy), RFs), (Rtuple, Rtuple')),
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
   163
      (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), names_lthy) = lthy
48975
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" activeAs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
      ||>> mk_Frees "b" activeAs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
      ||>> mk_Frees "b" activeBs
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
   168
      ||>> mk_Frees' "y" passiveAs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
      ||>> mk_Frees "B" BTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
      ||>> mk_Frees "B" BTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
      ||>> mk_Frees "B'" B'Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
      ||>> mk_Frees "B''" B''Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
      ||>> mk_Frees "s" sTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
      ||>> mk_Frees "sums" sum_sTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
      ||>> mk_Frees "s'" s'Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
      ||>> mk_Frees "s''" s''Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
      ||>> mk_Frees "f" fTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
      ||>> mk_Frees "f" fTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
      ||>> mk_Frees "f" self_fTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
      ||>> mk_Frees "g" gTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
      ||>> mk_Frees "g" all_gTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   182
      ||>> mk_Frees "x" FTsAs
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   183
      ||>> mk_Frees "y" FTsBs
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   184
      ||>> mk_Frees "y" FTsBs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
      ||>> mk_Frees "x" FRTs
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' "Rtuple") all_sbisT
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
      ||>> 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
   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" setRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
      ||>> mk_Frees "R'" setR'Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
      ||>> mk_Frees "R" setsRTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
      ||>> 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
   193
      ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT)
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
   194
      ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs);
48975
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;
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   197
    val passive_eqs = map HOLogic.eq_const passiveAs;
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;
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   214
    val map_comp0s = map map_comp0_of_bnf bnfs;
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   215
    val sym_map_comps = map mk_sym map_comp0s;
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   216
    val map_comps = 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;
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53258
diff changeset
   218
    val map_id0s = map map_id0_of_bnf bnfs;
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   219
    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
   220
    val set_bdss = map set_bd_of_bnf bnfs;
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   221
    val set_mapss = 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
   222
    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
   223
    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
   224
    val rel_Grps = map rel_Grp_of_bnf bnfs;
57726
18b8a8f10313 simplified tactics slightly
traytel
parents: 57700
diff changeset
   225
    val le_rel_OOs = map le_rel_OO_of_bnf bnfs;
56017
8d3df792d47e tuned tactic
traytel
parents: 56016
diff changeset
   226
    val in_rels = map in_rel_of_bnf bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
    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
   229
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
    (* derived thms *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
52956
1b62f05ab4fd honor user tfree names
blanchet
parents: 52938
diff changeset
   232
    (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
      map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   234
    fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
        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
   237
          (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
   238
        val rhs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
          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
   240
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   241
        Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (lhs, rhs))
55756
565a20b22f09 made tactics more robust
traytel
parents: 55644
diff changeset
   242
          (fn {context = ctxt, prems = _} => mk_map_comp_id_tac ctxt map_comp0)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   243
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   244
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   247
    val map_comp_id_thms = @{map 5} mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
    (*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
   250
      map id ... id f(m+1) ... f(m+n) x = x*)
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   251
    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
   252
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
        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
   254
          HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
            (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   256
        val prems = @{map 4} 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
   257
        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
   258
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   259
        Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, goal))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   260
          (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt 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
   261
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   262
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   265
    val map_cong0L_thms = @{map 5} mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
    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
   267
      (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
   268
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
    val map_arg_cong_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
      let
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   271
        val prems = map2 (curry mk_Trueprop_eq) yFs yFs_copy;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
   272
        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
   273
        val concls =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   274
          @{map 3} (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y))
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   275
            yFs yFs_copy maps;
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   276
        val goals = map2 (fn prem => fn concl => Logic.mk_implies (prem, concl)) prems concls;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   278
        map (fn goal =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   279
          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   280
            (hyp_subst_tac ctxt THEN' rtac ctxt refl) 1)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   281
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   282
          |> singleton (Proof_Context.export names_lthy lthy)) goals
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
    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
   286
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
    (* coalgebra *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   289
    val coalg_bind = mk_internal_b (coN ^ algN) ;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
    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
   291
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   292
    (*forall i = 1 ... n: (\<forall>x \<in> Bi. si \<in> Fi_in UNIV .. UNIV B1 ... Bn)*)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
    val coalg_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
      let
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   295
        val ins = @{map 3} mk_in (replicate n (passive_UNIVs @ Bs)) setssAs FTsAs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
        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
   297
          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
   298
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   299
        val rhs = Library.foldr1 HOLogic.mk_conj (@{map 5} mk_coalg_conjunct Bs ss ins zs zs')
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
      in
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   301
        fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) rhs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   302
      end;
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 ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
   305
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   306
      |> Local_Theory.define ((coalg_bind, NoSyn), (coalg_def_bind, coalg_spec))
49311
blanchet
parents: 49308
diff changeset
   307
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
   308
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   309
    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
   310
    val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free));
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   311
    val coalg_def = mk_unabs_def (2 * n) (Morphism.thm phi coalg_def_free RS meta_eq_to_obj_eq);
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   312
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   313
    fun mk_coalg Bs ss =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   314
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   315
        val args = Bs @ ss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
        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
   317
        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
   318
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   319
        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
   320
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   322
    val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
    val coalg_in_thms = map (fn i =>
52904
traytel
parents: 52839
diff changeset
   325
      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
   326
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
    val coalg_set_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   328
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   329
        val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss);
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
   330
        fun mk_prem x B = mk_Trueprop_mem (x, B);
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   331
        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
   332
        val prems = map2 mk_prem zs Bs;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   333
        val conclss = @{map 3} (fn s => fn x => fn sets => map2 (mk_concl s x) Bs (drop m sets))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   334
          ss zs setssAs;
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   335
        val goalss = map2 (fn prem => fn concls => map (fn concl =>
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   336
          Logic.list_implies (coalg_prem :: [prem], concl)) concls) prems conclss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   337
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   338
        map (fn goals => map (fn goal =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   339
          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   340
            mk_coalg_set_tac ctxt coalg_def)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   341
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   342
          |> singleton (Proof_Context.export names_lthy lthy)) goals) goalss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   345
    fun mk_tcoalg BTs = mk_coalg (map HOLogic.mk_UNIV BTs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   346
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   347
    val tcoalg_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
      let
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   349
        val goal = HOLogic.mk_Trueprop (mk_tcoalg activeAs ss)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   350
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   351
        Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   352
          (fn {context = ctxt, prems = _} => (rtac ctxt (coalg_def RS iffD2) 1 THEN CONJ_WRAP
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   353
            (K (EVERY' [rtac ctxt ballI, rtac ctxt CollectI,
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   354
              CONJ_WRAP' (K (EVERY' [rtac ctxt @{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
   355
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   356
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   358
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   359
    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
   360
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
    (* morphism *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   362
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   363
    val mor_bind = mk_internal_b morN;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   364
    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
   365
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   366
    (*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
   367
    (*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
   368
       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
   369
    val mor_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   370
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   371
        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
   372
          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
   373
        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
   374
          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
   375
            (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
   376
        val rhs = HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   377
          (Library.foldr1 HOLogic.mk_conj (@{map 5} mk_fbetw fs Bs B's zs zs'),
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   378
           Library.foldr1 HOLogic.mk_conj (@{map 7} mk_mor Bs mapsAsBs fs ss s's zs zs'))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
      in
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   380
        fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ fs) rhs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   381
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
    val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
   384
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   385
      |> Local_Theory.define ((mor_bind, NoSyn), (mor_def_bind, mor_spec))
49311
blanchet
parents: 49308
diff changeset
   386
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
   387
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   388
    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
   389
    val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   390
    val mor_def = mk_unabs_def (5 * n) (Morphism.thm phi mor_def_free RS meta_eq_to_obj_eq);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   391
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   392
    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
   393
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   394
        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
   395
        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
   396
        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
   397
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   398
        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
   399
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   400
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   401
    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
   402
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   403
    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
   404
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   405
        val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   406
        fun mk_image_goal f B1 B2 =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   407
          Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   408
        val image_goals = @{map 3} mk_image_goal fs Bs B's;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
        fun mk_elim_goal B mapAsBs f s s' x =
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
   410
          Logic.list_implies ([prem, mk_Trueprop_mem (x, B)],
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   411
            mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x)));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   412
        val elim_goals = @{map 6} mk_elim_goal Bs mapsAsBs fs ss s's zs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   413
        fun prove goal =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   414
          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   415
            mk_mor_elim_tac ctxt mor_def)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   416
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   417
          |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   419
        (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
   420
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   421
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
    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
   423
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   424
    val mor_incl_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   425
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   426
        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
   427
        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
   428
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   429
        Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   430
          (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   431
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   432
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   433
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   434
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   435
    val 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
   436
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   437
    val mor_comp_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   438
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   439
        val prems =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   440
          [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
   441
           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
   442
        val concl =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   443
          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
   444
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   445
        Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   446
          (K (mk_mor_comp_tac lthy 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
   447
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   448
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   449
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   450
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   451
    val mor_cong_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   452
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   453
        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
   454
         (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
   455
        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
   456
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   457
        Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58959
diff changeset
   458
          (K ((hyp_subst_tac lthy THEN' assume_tac lthy) 1))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   459
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   460
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   461
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   462
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   463
    val mor_UNIV_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   464
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   465
        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
   466
            (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
   467
            HOLogic.mk_comp (s', f));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   468
        val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   469
        val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   470
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   471
        Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (lhs, rhs))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   472
          (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt 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
   473
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   474
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   475
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   476
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   477
    val mor_str_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   479
        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
   480
          (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
   481
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   482
        Goal.prove_sorry lthy [] []
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   483
          (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   484
          (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt 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
   485
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   486
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   487
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   488
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
   489
    val mor_case_sum_thm =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   490
      let
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   491
        val maps = @{map 3} (fn s => fn sum_s => fn mapx =>
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
   492
          mk_case_sum (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
   493
          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
   494
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   495
        Goal.prove_sorry lthy [] []
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   496
          (HOLogic.mk_Trueprop (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   497
          (fn {context = ctxt, prems = _} => mk_mor_case_sum_tac ctxt 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
   498
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   499
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   500
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   501
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   502
    val 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
   503
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   504
    (* bisimulation *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   505
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   506
    val bis_bind = mk_internal_b bisN;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   507
    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
   508
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   509
    fun mk_bis_le_conjunct R B1 B2 = mk_leq R (mk_Times (B1, B2));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   510
    val bis_le = Library.foldr1 HOLogic.mk_conj (@{map 3} mk_bis_le_conjunct Rs Bs B's)
48975
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
    val bis_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   513
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   514
        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
   515
        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
   516
        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
   517
          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
   518
            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   519
            mk_Bex (mk_in (passive_UNIVs @ Rs) sets (snd (dest_Free RF)))
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   520
              (Term.absfree (dest_Free RF) (HOLogic.mk_conj
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   521
                (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
   522
                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
   523
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   524
        val rhs = HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   525
          (bis_le, Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   526
            (@{map 9} mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   527
      in
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   528
        fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ Rs) rhs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   529
      end;
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 ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
   532
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   533
      |> Local_Theory.define ((bis_bind, NoSyn), (bis_def_bind, bis_spec))
49311
blanchet
parents: 49308
diff changeset
   534
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
   535
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   536
    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
   537
    val bis = fst (Term.dest_Const (Morphism.term phi bis_free));
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   538
    val bis_def = mk_unabs_def (5 * n) (Morphism.thm phi bis_def_free RS meta_eq_to_obj_eq);
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   539
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   540
    fun mk_bis Bs1 ss1 Bs2 ss2 Rs =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   541
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   542
        val args = Bs1 @ ss1 @ Bs2 @ ss2 @ Rs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   543
        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
   544
        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
   545
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   546
        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
   547
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   548
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   549
    val bis_cong_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   550
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   551
        val prems = map HOLogic.mk_Trueprop
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   552
         (mk_bis Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs)
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   553
        val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs_copy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   554
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   555
        Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58959
diff changeset
   556
          (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   557
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   558
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   559
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   560
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   561
    val bis_rel_thm =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   562
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
   563
        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
   564
          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
   565
            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   566
            Term.list_comb (rel, passive_eqs @ map mk_in_rel 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
   567
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   568
        val rhs = HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   569
          (bis_le, Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   570
            (@{map 6} mk_conjunct Rs ss s's zs z's relsAsBs))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   571
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   572
        Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (mk_bis Bs ss B's s's Rs, rhs))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   573
          (fn {context = ctxt, prems = _} => mk_bis_rel_tac ctxt m bis_def in_rels map_comps
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   574
            map_cong0s set_mapss)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   575
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   576
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   577
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   578
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   579
    val bis_converse_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   580
      Goal.prove_sorry lthy [] []
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   581
        (Logic.mk_implies (HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs),
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   582
          HOLogic.mk_Trueprop (mk_bis B's s's Bs ss (map mk_converse Rs))))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   583
        (fn {context = ctxt, prems = _} => mk_bis_converse_tac ctxt m bis_rel_thm rel_congs
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   584
          rel_converseps)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   585
      |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   586
      |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   587
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   588
    val bis_O_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   589
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   590
        val prems =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   591
          [HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs),
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   592
           HOLogic.mk_Trueprop (mk_bis B's s's B''s s''s R's)];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   593
        val concl =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   594
          HOLogic.mk_Trueprop (mk_bis Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   595
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   596
        Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
57726
18b8a8f10313 simplified tactics slightly
traytel
parents: 57700
diff changeset
   597
          (K (mk_bis_O_tac lthy m bis_rel_thm rel_congs le_rel_OOs))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   598
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   599
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   600
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   601
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   602
    val bis_Gr_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   603
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   604
        val concl =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   605
          HOLogic.mk_Trueprop (mk_bis Bs ss B's s's (map2 mk_Gr Bs fs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   606
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   607
        Goal.prove_sorry lthy [] [] (Logic.list_implies ([coalg_prem, mor_prem], concl))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
   608
          (fn {context = ctxt, prems = _} => mk_bis_Gr_tac ctxt bis_rel_thm rel_Grps mor_image_thms
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
   609
            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
   610
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   611
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   612
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   613
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   614
    val 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
   615
      ((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
   616
      replicate n @{thm image2_Gr});
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   617
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
   618
    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
   619
      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
   620
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   621
    val bis_Union_thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   622
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   623
        val prem =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   624
          HOLogic.mk_Trueprop (mk_Ball Idx
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   625
            (Term.absfree idx' (mk_bis Bs ss B's s's (map (fn R => R $ idx) Ris))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   626
        val concl =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   627
          HOLogic.mk_Trueprop (mk_bis Bs ss B's s's (map (mk_UNION Idx) Ris));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   628
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   629
        Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, concl))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
   630
          (fn {context = ctxt, prems = _} => mk_bis_Union_tac ctxt 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
   631
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   632
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   633
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   634
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   635
    (* self-bisimulation *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   636
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   637
    fun mk_sbis Bs ss Rs = mk_bis Bs ss Bs ss Rs;
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   638
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   639
    val sbis_prem = HOLogic.mk_Trueprop (mk_sbis Bs ss sRs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   640
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   641
    (* largest self-bisimulation *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   642
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   643
    val lsbis_binds = mk_internal_bs lsbisN;
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   644
    fun lsbis_bind i = nth lsbis_binds (i - 1);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   645
    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
   646
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   647
    val all_sbis = HOLogic.mk_Collect (fst Rtuple', snd Rtuple', list_exists_free sRs
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   648
      (HOLogic.mk_conj (HOLogic.mk_eq (Rtuple, HOLogic.mk_tuple sRs), mk_sbis Bs ss sRs)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   649
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   650
    fun lsbis_spec i =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   651
      fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss)
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   652
        (mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   653
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   654
    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
   655
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   656
      |> fold_map (fn i => Local_Theory.define
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   657
        ((lsbis_bind i, NoSyn), (lsbis_def_bind i, lsbis_spec i))) ks
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   658
      |>> 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
   659
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   660
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   661
    val 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
   662
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   663
    val lsbis_defs = map (fn def =>
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   664
      mk_unabs_def (2 * n) (Morphism.thm phi def RS meta_eq_to_obj_eq)) lsbis_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   665
    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
   666
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   667
    fun mk_lsbis Bs ss i =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   668
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   669
        val args = Bs @ ss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   670
        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
   671
        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
   672
        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
   673
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   674
        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
   675
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   676
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   677
    val sbis_lsbis_thm =
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   678
      Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_sbis Bs ss (map (mk_lsbis Bs ss) ks)))
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
   679
        (K (mk_sbis_lsbis_tac lthy lsbis_defs bis_Union_thm bis_cong_thm))
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   680
      |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   681
      |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   682
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   683
    val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS
52904
traytel
parents: 52839
diff changeset
   684
      (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
   685
    val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS
52904
traytel
parents: 52839
diff changeset
   686
      (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
   687
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   688
    val incl_lsbis_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   689
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   690
        fun mk_concl i R = HOLogic.mk_Trueprop (mk_leq R (mk_lsbis Bs ss i));
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   691
        val goals = map2 (fn i => fn R => Logic.mk_implies (sbis_prem, mk_concl i R)) ks sRs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   692
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   693
        @{map 3} (fn goal => fn i => fn def =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   694
          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   695
            mk_incl_lsbis_tac ctxt n i def)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   696
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   697
          |> singleton (Proof_Context.export names_lthy lthy)) goals ks lsbis_defs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   698
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   699
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   700
    val equiv_lsbis_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   701
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   702
        fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis Bs ss i));
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   703
        val goals = map2 (fn i => fn B => Logic.mk_implies (coalg_prem, mk_concl i B)) ks Bs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   704
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   705
        @{map 3} (fn goal => fn l_incl => fn incl_l =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
   706
          Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   707
            (fn {context = ctxt, prems = _} => mk_equiv_lsbis_tac ctxt sbis_lsbis_thm l_incl incl_l
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   708
              bis_Id_on_thm bis_converse_thm bis_O_thm)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   709
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
   710
          |> singleton (Proof_Context.export names_lthy lthy))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   711
        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
   712
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   713
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   714
    val timer = time (timer "Bisimulations");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   715
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   716
    (* bounds *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   717
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   718
    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
   719
      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
   720
      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
   721
      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
   722
      else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   723
        let
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   724
          val sbdT_bind = mk_internal_b sum_bdTN;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   725
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   726
          val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =
56516
a13c2ccc160b more accurate type arguments for intermeadiate typedefs
traytel
parents: 56348
diff changeset
   727
            typedef (sbdT_bind, sum_bdT_params', NoSyn)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   728
              (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
   729
                EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   730
56516
a13c2ccc160b more accurate type arguments for intermeadiate typedefs
traytel
parents: 56348
diff changeset
   731
          val sbdT = Type (sbdT_name, sum_bdT_params);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   732
          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
   733
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   734
          val sbd_bind = mk_internal_b sum_bdN;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   735
          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
   736
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   737
          val sbd_spec = mk_dir_image sum_bd Abs_sbdT;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   738
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   739
          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
   740
            lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   741
            |> Local_Theory.define ((sbd_bind, NoSyn), (sbd_def_bind, sbd_spec))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   742
            ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   743
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   744
          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
   745
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   746
          val sbd_def = Morphism.thm phi sbd_def_free RS meta_eq_to_obj_eq;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   747
          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
   748
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49225
diff changeset
   749
          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
   750
          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
   751
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   752
          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
   753
          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
   754
          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
   755
55770
f2cf7f92c9ac intermediate typedef for the type of the bound (local to lfp)
traytel
parents: 55756
diff changeset
   756
          val sbd_ordIso = @{thm ssubst_Pair_rhs} OF
f2cf7f92c9ac intermediate typedef for the type of the bound (local to lfp)
traytel
parents: 55756
diff changeset
   757
            [@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order], sbd_def];
f2cf7f92c9ac intermediate typedef for the type of the bound (local to lfp)
traytel
parents: 55756
diff changeset
   758
          val sbd_card_order = @{thm iffD2[OF arg_cong[of _ _ card_order]]} OF
f2cf7f92c9ac intermediate typedef for the type of the bound (local to lfp)
traytel
parents: 55756
diff changeset
   759
            [sbd_def, @{thm card_order_dir_image} OF [Abs_sbdT_bij, sum_card_order]];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   760
          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
   761
          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
   762
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   763
          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
   764
            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
   765
              [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   766
          val set_sbdss = @{map 3} mk_set_sbd ks bd_Card_orders set_bdss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   767
       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
   768
         (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
   769
       end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   770
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   771
    val sbdTs = replicate n sbdT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   772
    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
   773
    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
   774
    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
   775
    val bdTs = passiveAs @ replicate n sbdT;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   776
    val to_sbd_maps = @{map 4} mk_map_of_bnf Dss Ass (replicate n bdTs) bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   777
    val bdFTs = mk_FTs bdTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   778
    val sbdFT = mk_sumTN bdFTs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   779
    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
   780
    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
   781
    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
   782
    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
   783
    val treeFTs = mk_FTs treeTs;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   784
    val tree_maps = @{map 4} mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   785
    val final_maps = @{map 4} mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   786
    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
   787
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   788
    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
   789
    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
   790
    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
   791
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   792
    val Nil = HOLogic.mk_tuple (@{map 3} (fn i => fn z => fn z'=>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   793
      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
   794
    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
   795
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   796
    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
   797
      (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
   798
      names_lthy) = names_lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   799
      |> 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
   800
      ||>> mk_Frees' "k" sbdTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   801
      ||>> 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
   802
      ||>> 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
   803
      ||>> 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
   804
      ||>> 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
   805
      ||>> 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
   806
      ||>> mk_Frees "x" bdFTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   807
      ||>> 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
   808
      ||>> 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
   809
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   810
    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
   811
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   812
    val timer = time (timer "Bounds");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   813
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   814
    (* tree coalgebra *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   815
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   816
    val isNode_binds = mk_internal_bs isNodeN;
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   817
    fun isNode_bind i = nth isNode_binds (i - 1);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   818
    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
   819
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   820
    val isNodeT =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   821
      Library.foldr (op -->) (map fastype_of [Kl, lab, kl], HOLogic.boolT);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   822
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   823
    val Succs = @{map 3} (fn i => fn k => fn k' =>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   824
      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
   825
      ks kks kks';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   826
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   827
    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
   828
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   829
        val active_sets = drop m (map (fn set => set $ x) sets);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   830
        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
   831
          (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) ::
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   832
          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
   833
      in
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   834
        fold_rev (Term.absfree o Term.dest_Free) [Kl, lab, kl] rhs
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 ((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
   838
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   839
      |> @{fold_map 3} (fn i => fn x => fn sets => Local_Theory.define
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   840
        ((isNode_bind i, NoSyn), (isNode_def_bind i, isNode_spec sets x i)))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   841
        ks xs isNode_setss
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   842
      |>> 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
   843
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   844
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   845
    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
   846
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   847
    val isNode_defs = map (fn def =>
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   848
      mk_unabs_def 3 (Morphism.thm phi def RS meta_eq_to_obj_eq)) isNode_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   849
    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
   850
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   851
    fun mk_isNode kl i =
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   852
      Term.list_comb (Const (nth isNodes (i - 1), isNodeT), [Kl, lab, kl]);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   853
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   854
    val isTree =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   855
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   856
        val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   857
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   858
        val tree = mk_Ball Kl (Term.absfree kl'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   859
          (Library.foldr1 HOLogic.mk_conj (@{map 4} (fn Succ => fn i => fn k => fn k' =>
55581
d1c228753d76 another simplification of internal codatatype construction
traytel
parents: 55577
diff changeset
   860
            mk_Ball Succ (Term.absfree k' (mk_isNode
d1c228753d76 another simplification of internal codatatype construction
traytel
parents: 55577
diff changeset
   861
              (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i)))
d1c228753d76 another simplification of internal codatatype construction
traytel
parents: 55577
diff changeset
   862
          Succs ks kks kks')));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   863
      in
55577
a6c2379078c8 simplifications of internal codatatype construction
traytel
parents: 55541
diff changeset
   864
        HOLogic.mk_conj (empty, tree)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   865
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   866
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   867
    val carT_binds = mk_internal_bs carTN;
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   868
    fun carT_bind i = nth carT_binds (i - 1);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   869
    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
   870
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   871
    fun carT_spec i =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   872
      HOLogic.mk_Collect (fst Kl_lab', snd Kl_lab', list_exists_free [Kl, lab]
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   873
        (HOLogic.mk_conj (HOLogic.mk_eq (Kl_lab, HOLogic.mk_prod (Kl, lab)),
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   874
          HOLogic.mk_conj (isTree, mk_isNode (HOLogic.mk_list sum_sbdT []) i))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   875
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   876
    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
   877
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   878
      |> fold_map (fn i =>
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   879
        Local_Theory.define ((carT_bind i, NoSyn), (carT_def_bind i, carT_spec i))) ks
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   880
      |>> 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
   881
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   882
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   883
    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
   884
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   885
    val carT_defs = map (fn def => Morphism.thm phi def RS meta_eq_to_obj_eq) carT_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   886
    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
   887
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   888
    fun mk_carT i = Const (nth carTs (i - 1), HOLogic.mk_setT treeT);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   889
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   890
    val strT_binds = mk_internal_bs strTN;
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   891
    fun strT_bind i = nth strT_binds (i - 1);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   892
    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
   893
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   894
    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
   895
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   896
        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
   897
          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
   898
          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
   899
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   900
        val f = Term.list_comb (mapFT, passive_ids @ @{map 3} mk_f ks kks kks');
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   901
        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
   902
        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
   903
      in
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   904
        HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
   905
          (mk_case_sumN fs $ (lab $ HOLogic.mk_list sum_sbdT []))))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   906
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   907
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   908
    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
   909
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   910
      |> @{fold_map 3} (fn i => fn mapFT => fn FT => Local_Theory.define
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   911
        ((strT_bind i, NoSyn), (strT_def_bind i, strT_spec mapFT FT i)))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   912
        ks tree_maps treeFTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   913
      |>> 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
   914
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   915
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   916
    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
   917
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   918
    val strT_defs = map (fn def =>
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55602
diff changeset
   919
        trans OF [Morphism.thm phi def RS meta_eq_to_obj_eq RS fun_cong, @{thm prod.case}])
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   920
      strT_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   921
    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
   922
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   923
    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
   924
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   925
    val carTAs = map mk_carT ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   926
    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
   927
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   928
    val coalgT_thm =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
   929
      Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_coalg carTAs strTAs))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
   930
        (fn {context = ctxt, prems = _} => mk_coalgT_tac ctxt m
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
   931
          (coalg_def :: isNode_defs @ carT_defs) strT_defs set_mapss)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
   932
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   933
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   934
    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
   935
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   936
    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
   937
      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
   938
    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
   939
      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
   940
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   941
    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
   942
      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
   943
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   944
    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
   945
    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
   946
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   947
    val Lev_bind = mk_internal_b LevN;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   948
    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
   949
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   950
    val Lev_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   951
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   952
        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
   953
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   954
            val sets = drop m setsAs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   955
            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
   956
              let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   957
                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
   958
                  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
   959
                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
   960
                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
   961
              in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   962
                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
   963
                  (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
   964
              end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   965
          in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   966
            Term.absfree a' (Library.foldl1 mk_union (@{map 3} mk_set ks sets zs_copy))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   967
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   968
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   969
        val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
   970
          (HOLogic.mk_tuple (@{map 5} mk_Suc ks ss setssAs zs zs')));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   971
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   972
        val rhs = mk_rec_nat Zero Suc;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   973
      in
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   974
        fold_rev (Term.absfree o Term.dest_Free) ss rhs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   975
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   976
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   977
    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
   978
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   979
      |> Local_Theory.define ((Lev_bind, NoSyn), (Lev_def_bind, Lev_spec))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   980
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   981
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   982
    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
   983
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
   984
    val Lev_def = mk_unabs_def n (Morphism.thm phi Lev_def_free RS meta_eq_to_obj_eq);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   985
    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
   986
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   987
    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
   988
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   989
        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
   990
        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
   991
          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
   992
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   993
        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
   994
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   995
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   996
    val Lev_0s = flat (mk_rec_simps n @{thm rec_nat_0_imp} [Lev_def]);
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   997
    val Lev_Sucs = flat (mk_rec_simps n @{thm rec_nat_Suc_imp} [Lev_def]);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   998
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
   999
    val rv_bind = mk_internal_b rvN;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1000
    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
  1001
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1002
    val rv_spec =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1003
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1004
        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
  1005
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1006
            fun mk_case i' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1007
              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
  1008
          in
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1009
            Term.absfree b' (mk_case_sumN (map mk_case ks) $ sumx)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1010
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1011
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1012
        val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1013
          (HOLogic.mk_tuple (@{map 4} mk_Cons ks ss zs zs'))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1014
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55393
diff changeset
  1015
        val rhs = mk_rec_list Nil Cons;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1016
      in
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1017
        fold_rev (Term.absfree o Term.dest_Free) ss rhs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1018
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1019
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1020
    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
  1021
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1022
      |> Local_Theory.define ((rv_bind, NoSyn), (rv_def_bind, rv_spec))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1023
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1024
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1025
    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
  1026
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1027
    val rv_def = mk_unabs_def n (Morphism.thm phi rv_def_free RS meta_eq_to_obj_eq);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1028
    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
  1029
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1030
    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
  1031
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1032
        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
  1033
        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
  1034
        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
  1035
          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
  1036
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1037
        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
  1038
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1039
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55393
diff changeset
  1040
    val rv_Nils = flat (mk_rec_simps n @{thm rec_list_Nil_imp} [rv_def]);
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55393
diff changeset
  1041
    val rv_Conss = flat (mk_rec_simps n @{thm rec_list_Cons_imp} [rv_def]);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1042
53566
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
  1043
    val beh_binds = mk_internal_bs behN;
5ff3a2d112d7 conceal internal bindings
traytel
parents: 53553
diff changeset
  1044
    fun beh_bind i = nth beh_binds (i - 1);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1045
    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
  1046
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1047
    fun beh_spec i z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1048
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1049
        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
  1050
          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
  1051
            (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
  1052
55577
a6c2379078c8 simplifications of internal codatatype construction
traytel
parents: 55541
diff changeset
  1053
        val Lab = Term.absfree kl'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1054
          (mk_case_sumN (@{map 5} mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1055
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1056
        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
  1057
          (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
  1058
      in
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1059
        fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) rhs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1060
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1061
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1062
    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
  1063
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1064
      |> @{fold_map 2} (fn i => fn z =>
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1065
        Local_Theory.define ((beh_bind i, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1066
      |>> 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
  1067
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1068
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1069
    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
  1070
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1071
    val beh_defs = map (fn def =>
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1072
      mk_unabs_def (n + 1) (Morphism.thm phi def RS meta_eq_to_obj_eq)) beh_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1073
    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
  1074
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1075
    fun mk_beh ss i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1076
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1077
        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
  1078
        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
  1079
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1080
        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
  1081
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1082
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1083
    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
  1084
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1085
        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
  1086
          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
  1087
        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
  1088
          (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
  1089
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1090
        val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1091
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1092
        val length_Lev =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1093
          Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1094
            (K (mk_length_Lev_tac lthy cts Lev_0s Lev_Sucs))
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1095
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1096
          |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1097
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1098
        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
  1099
        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
  1100
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1101
        fun mk_goal i z = Logic.mk_implies
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1102
            (mk_Trueprop_mem (kl, mk_Lev ss nat i $ z),
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1103
            mk_Trueprop_mem (kl, mk_Lev ss (mk_size kl) i $ z));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1104
        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
  1105
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1106
        val length_Levs' = map2 (fn goal => fn length_Lev =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1107
          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1108
            mk_length_Lev'_tac ctxt length_Lev)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1109
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1110
          |> singleton (Proof_Context.export names_lthy lthy)) goals length_Levs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1111
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1112
        (length_Levs, length_Levs')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1113
      end;
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 rv_last_thmss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1116
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1117
        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
  1118
          (HOLogic.mk_eq
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1119
            (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
  1120
            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
  1121
        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
  1122
          (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
  1123
            Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1124
              (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
  1125
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1126
        val cTs = [SOME (Thm.ctyp_of lthy sum_sbdT)];
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1127
        val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree kl' goal, kl];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1128
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1129
        val rv_last =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1130
          Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1131
            (fn {context = ctxt, prems = _} => mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1132
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1133
          |> singleton (Proof_Context.export names_lthy lthy);
48975
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
        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
  1136
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1137
        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
  1138
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1139
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1140
    val set_Lev_thmsss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1141
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1142
        fun mk_conjunct i z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1143
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1144
            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
  1145
              let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1146
                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
  1147
                  (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
  1148
                    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
  1149
                      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
  1150
                      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
  1151
              in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1152
                HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z' i'),
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1153
                  (Library.foldr1 HOLogic.mk_conj
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1154
                    (@{map 3} mk_conjunct'' ks (drop m sets) zs_copy2)))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1155
              end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1156
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1157
            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1158
              Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct' ks setssAs ss zs_copy))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1159
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1160
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1161
        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
  1162
          (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
  1163
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1164
        val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1165
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1166
        val set_Lev =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1167
          Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1168
            (K (mk_set_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss))
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1169
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1170
          |> singleton (Proof_Context.export names_lthy lthy);
48975
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
        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
  1173
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1174
        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
  1175
          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
  1176
          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
  1177
          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
  1178
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1179
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1180
    val set_image_Lev_thmsss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1181
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1182
        fun mk_conjunct i z =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1183
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1184
            fun mk_conjunct' i' sets =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1185
              let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1186
                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
  1187
                  (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
  1188
                  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
  1189
              in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1190
                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
  1191
                  (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
  1192
                    mk_Lev ss (HOLogic.mk_Suc nat) i $ z),
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1193
                  (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct'' ks sets ss zs_copy)))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1194
              end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1195
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1196
            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
  1197
              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
  1198
          end;
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
        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
  1201
          (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
  1202
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1203
        val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1204
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1205
        val set_image_Lev =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1206
          Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51782
diff changeset
  1207
            (K (mk_set_image_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1208
              from_to_sbd_thmss to_sbd_inj_thmss))
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1209
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1210
          |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1211
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1212
        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
  1213
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1214
        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
  1215
          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
  1216
          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
  1217
          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
  1218
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1219
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1220
    val mor_beh_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1221
      Goal.prove_sorry lthy [] []
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1222
        (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss carTAs strTAs (map (mk_beh ss) ks)))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1223
        (fn {context = ctxt, prems = _} => mk_mor_beh_tac ctxt m mor_def mor_cong_thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1224
          beh_defs carT_defs strT_defs isNode_defs
55577
a6c2379078c8 simplifications of internal codatatype construction
traytel
parents: 55541
diff changeset
  1225
          to_sbd_inj_thmss from_to_sbd_thmss Lev_0s Lev_Sucs rv_Nils rv_Conss
a6c2379078c8 simplifications of internal codatatype construction
traytel
parents: 55541
diff changeset
  1226
          length_Lev_thms length_Lev'_thms rv_last_thmss set_Lev_thmsss
a6c2379078c8 simplifications of internal codatatype construction
traytel
parents: 55541
diff changeset
  1227
          set_image_Lev_thmsss set_mapss map_comp_id_thms map_cong0s)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1228
      |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1229
      |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1230
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1231
    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
  1232
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1233
    val lsbisAs = map (mk_lsbis carTAs strTAs) ks;
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1234
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1235
    fun mk_str_final i =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1236
      mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1),
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1237
        passive_ids @ map mk_proj lsbisAs), nth strTAs (i - 1)));
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1238
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1239
    val car_finals = map2 mk_quotient carTAs lsbisAs;
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1240
    val str_finals = map mk_str_final ks;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1241
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1242
    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
  1243
    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
  1244
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1245
    val congruent_str_final_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1246
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1247
        fun mk_goal R final_map strT =
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1248
          HOLogic.mk_Trueprop (mk_congruent R (HOLogic.mk_comp
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1249
            (Term.list_comb (final_map, passive_ids @ map mk_proj lsbisAs), strT)));
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1250
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1251
        val goals = @{map 3} mk_goal lsbisAs final_maps strTAs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1252
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1253
        @{map 4} (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1254
          Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1255
            (fn {context = ctxt, prems = _} => mk_congruent_str_final_tac ctxt m lsbisE map_comp_id
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1256
              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
  1257
          |> Thm.close_derivation)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1258
        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
  1259
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1260
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1261
    val coalg_final_thm = Goal.prove_sorry lthy [] []
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1262
      (HOLogic.mk_Trueprop (mk_coalg car_finals str_finals))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1263
      (fn {context = ctxt, prems = _} => mk_coalg_final_tac ctxt m coalg_def
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1264
        congruent_str_final_thms equiv_LSBIS_thms set_mapss coalgT_set_thmss)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1265
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1266
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1267
    val mor_T_final_thm = Goal.prove_sorry lthy [] []
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1268
      (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finals str_finals (map mk_proj lsbisAs)))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1269
      (fn {context = ctxt, prems = _} => mk_mor_T_final_tac ctxt mor_def congruent_str_final_thms
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1270
        equiv_LSBIS_thms)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1271
      |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1272
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1273
    val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm];
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1274
    val in_car_final_thms = map (fn thm => thm OF [mor_final_thm, UNIV_I]) mor_image'_thms;
48975
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 timer = time (timer "Final coalgebra");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1277
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1278
    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
  1279
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1280
      |> @{fold_map 4} (fn b => fn mx => fn car_final => fn in_car_final =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1281
          typedef (b, params, mx) car_final NONE
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1282
            (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt in_car_final] 1))
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1283
        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
  1284
      |>> 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
  1285
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1286
    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
  1287
    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
  1288
    val Ts' = mk_Ts passiveBs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1289
    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
  1290
    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
  1291
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1292
    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
  1293
    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
  1294
    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
  1295
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1296
    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
  1297
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1298
    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
  1299
    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
  1300
    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
  1301
    val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1302
    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
  1303
    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
  1304
    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
  1305
    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
  1306
    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
  1307
    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
  1308
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1309
    val emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1310
    val Zeros = map (fn empty =>
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1311
     HOLogic.mk_tuple (map (fn U => absdummy U empty) Ts)) emptys;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1312
    val hrecTs = map fastype_of Zeros;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1313
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1314
    val (((((((((((((Jzs, Jzs'), Jz's), Jzs_copy), Jz's_copy), Jzs1), Jzs2),
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1315
      TRs), unfold_fs), corec_ss), phis), dtor_set_induct_phiss), (hrecs, hrecs')),
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1316
      names_lthy) = names_lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1317
      |> mk_Frees' "z" Ts
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1318
      ||>> mk_Frees "y" Ts'
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1319
      ||>> mk_Frees "z'" Ts
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1320
      ||>> mk_Frees "y'" Ts'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1321
      ||>> mk_Frees "z1" Ts
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1322
      ||>> mk_Frees "z2" Ts
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49588
diff changeset
  1323
      ||>> mk_Frees "r" (map (mk_relT o `I) Ts)
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1324
      ||>> mk_Frees "f" unfold_fTs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1325
      ||>> mk_Frees "s" corec_sTs
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1326
      ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts)
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1327
      ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs)
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1328
      ||>> mk_Frees' "rec" hrecTs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1329
54492
6fae4ecd4ab3 prefix internal names as well
blanchet
parents: 54421
diff changeset
  1330
    fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59819
diff changeset
  1331
    val dtor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o dtor_bind;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1332
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1333
    fun dtor_spec rep str map_FT Jz Jz' =
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1334
      Term.absfree Jz'
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1335
        (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $ (str $ (rep $ Jz)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1336
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1337
    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
  1338
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1339
      |> @{fold_map 6} (fn i => fn rep => fn str => fn mapx => fn Jz => fn Jz' =>
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1340
        Local_Theory.define ((dtor_bind i, NoSyn),
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1341
          (dtor_def_bind i, dtor_spec rep str mapx Jz Jz')))
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1342
        ks Rep_Ts str_finals map_FTs Jzs Jzs'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1343
      |>> 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
  1344
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1345
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1346
    val phi = Proof_Context.export_morphism lthy_old lthy;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1347
    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
  1348
      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
  1349
        Morphism.term phi) dtor_frees;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1350
    val dtors = mk_dtors passiveAs;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1351
    val dtor's = mk_dtors passiveBs;
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1352
    val dtor_defs = map (fn def =>
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1353
      Morphism.thm phi def RS meta_eq_to_obj_eq RS fun_cong) dtor_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1354
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1355
    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
  1356
    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
  1357
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1358
        val mor_Rep =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1359
          Goal.prove_sorry lthy [] []
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1360
            (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts))
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1361
            (fn {context = ctxt, prems = _} => mk_mor_Rep_tac ctxt (mor_def :: dtor_defs) Reps
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1362
              Abs_inverses coalg_final_set_thmss 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
  1363
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1364
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1365
        val mor_Abs =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1366
          Goal.prove_sorry lthy [] []
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1367
            (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1368
            (fn {context = ctxt, prems = _} => mk_mor_Abs_tac ctxt (mor_def :: dtor_defs)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1369
              Abs_inverses)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  1370
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1371
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1372
        (mor_Rep, mor_Abs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1373
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1374
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1375
    val timer = time (timer "dtor definitions & thms");
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1376
54492
6fae4ecd4ab3 prefix internal names as well
blanchet
parents: 54421
diff changeset
  1377
    fun unfold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_unfoldN ^ "_");
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59819
diff changeset
  1378
    val unfold_def_bind = rpair [] o Binding.concealed o Thm.def_binding o unfold_bind;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1379
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1380
    fun unfold_spec abs f z = fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) (abs $ (f $ z));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1381
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1382
    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
  1383
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1384
      |> @{fold_map 4} (fn i => fn abs => fn f => fn z =>
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1385
        Local_Theory.define ((unfold_bind i, NoSyn), (unfold_def_bind i, unfold_spec abs f z)))
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1386
          ks Abs_Ts (map (fn i => HOLogic.mk_comp
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1387
            (mk_proj (nth lsbisAs (i - 1)), mk_beh ss i)) ks) zs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1388
      |>> 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
  1389
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1390
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1391
    val phi = Proof_Context.export_morphism lthy_old lthy;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1392
    val unfolds = map (Morphism.term phi) unfold_frees;
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1393
    val unfold_names = map (fst o dest_Const) unfolds;
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1394
    fun mk_unfolds passives actives =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1395
      @{map 3} (fn name => fn T => fn active =>
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1396
        Const (name, Library.foldr (op -->)
52923
traytel
parents: 52913
diff changeset
  1397
          (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
  1398
      unfold_names (mk_Ts passives) actives;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1399
    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
  1400
      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1401
    val unfold_defs = map (fn def =>
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1402
      mk_unabs_def (n + 1) (Morphism.thm phi def RS meta_eq_to_obj_eq)) unfold_def_frees;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1403
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1404
    val mor_unfold_thm =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1405
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1406
        val Abs_inverses' = map2 (curry op RS) in_car_final_thms Abs_inverses;
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1407
        val morEs' = map (fn thm => (thm OF [mor_final_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
  1408
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1409
        Goal.prove_sorry lthy [] []
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1410
          (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks)))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1411
          (fn {context = ctxt, prems = _} => mk_mor_unfold_tac ctxt m mor_UNIV_thm dtor_defs
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1412
            unfold_defs Abs_inverses' morEs' 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
  1413
        |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1414
        |> singleton (Proof_Context.export names_lthy lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1415
      end;
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1416
    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
  1417
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1418
    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
  1419
      let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1420
        val prem = HOLogic.mk_Trueprop (mk_sbis UNIVs dtors TRs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1421
        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
  1422
          (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
  1423
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1424
        `split_conj_thm (Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, concl))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1425
          (fn {context = ctxt, prems = _} => mk_raw_coind_tac ctxt bis_def bis_cong_thm bis_O_thm
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1426
            bis_converse_thm bis_Gr_thm tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1427
            lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1428
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1429
          |> singleton (Proof_Context.export names_lthy lthy))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1430
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1431
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1432
    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
  1433
      let
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1434
        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
  1435
        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
  1436
        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
  1437
          (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
  1438
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1439
        val bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm);
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1440
        val mor_thm = mor_comp_thm OF [mor_final_thm, mor_Abs_thm];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1441
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1442
        val unique_mor = Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, unique))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1443
          (fn {context = ctxt, prems = _} => mk_unfold_unique_mor_tac ctxt raw_coind_thms
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1444
            bis_thm mor_thm unfold_defs)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1445
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1446
          |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1447
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1448
        `split_conj_thm unique_mor
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1449
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1450
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1451
    val (dtor_unfold_unique_thms, dtor_unfold_unique_thm) = `split_conj_thm (split_conj_prems n
52904
traytel
parents: 52839
diff changeset
  1452
      (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
  1453
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1454
    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
  1455
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1456
    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
  1457
      let
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1458
        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
  1459
      in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1460
        map2 (fn unique => fn unfold_ctor =>
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1461
          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
  1462
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1463
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1464
    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
  1465
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1466
    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
  1467
      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
  1468
        map HOLogic.id_const passiveAs @ dtors)) Dss bnfs;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1469
54492
6fae4ecd4ab3 prefix internal names as well
blanchet
parents: 54421
diff changeset
  1470
    fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59819
diff changeset
  1471
    val ctor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o ctor_bind;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1472
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1473
    fun ctor_spec i = 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
  1474
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1475
    val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
  1476
      lthy
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1477
      |> fold_map (fn i =>
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1478
        Local_Theory.define ((ctor_bind i, NoSyn), (ctor_def_bind i, ctor_spec i))) ks
49311
blanchet
parents: 49308
diff changeset
  1479
      |>> apsnd split_list o split_list
blanchet
parents: 49308
diff changeset
  1480
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
  1481
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1482
    val phi = Proof_Context.export_morphism lthy_old lthy;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1483
    fun mk_ctors params =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1484
      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
  1485
        ctor_frees;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1486
    val ctors = mk_ctors params';
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1487
    val ctor_defs = map (fn def => Morphism.thm phi def RS meta_eq_to_obj_eq) ctor_def_frees;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1488
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1489
    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
  1490
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1491
    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
  1492
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1493
        fun mk_goal dtor ctor FT =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1494
         mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1495
        val goals = @{map 3} 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
  1496
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1497
        @{map 5} (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
  1498
          Goal.prove_sorry lthy [] [] goal
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1499
            (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt ctor_def unfold map_comp_id
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1500
              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
  1501
          |> Thm.close_derivation)
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1502
          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
  1503
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1504
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1505
    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
  1506
    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
  1507
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1508
    val bij_dtor_thms =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1509
      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
  1510
    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
  1511
    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
  1512
    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
  1513
    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
  1514
    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
  1515
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1516
    val bij_ctor_thms =
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1517
      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
  1518
    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
  1519
    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
  1520
    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
  1521
    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
  1522
    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
  1523
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1524
    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
  1525
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1526
    val corec_Inl_sum_thms =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1527
      let
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1528
        val mor = mor_comp_thm OF [mor_case_sum_thm, mor_unfold_thm];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1529
      in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1530
        map2 (fn unique => fn unfold_dtor =>
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
  1531
          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
  1532
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1533
54492
6fae4ecd4ab3 prefix internal names as well
blanchet
parents: 54421
diff changeset
  1534
    fun corec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_corecN ^ "_");
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59819
diff changeset
  1535
    val corec_def_bind = rpair [] o Binding.concealed o Thm.def_binding o corec_bind;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1536
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1537
    val corec_strs =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1538
      @{map 3} (fn dtor => fn sum_s => fn mapx =>
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1539
        mk_case_sum
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1540
          (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
  1541
      dtors corec_ss corec_maps;
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1542
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1543
    fun corec_spec i T AT =
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1544
      fold_rev (Term.absfree o Term.dest_Free) corec_ss
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1545
        (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
  1546
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1547
    val ((corec_frees, (_, corec_def_frees)), (lthy, lthy_old)) =
49311
blanchet
parents: 49308
diff changeset
  1548
      lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1549
      |> @{fold_map 3} (fn i => fn T => fn AT =>
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1550
        Local_Theory.define ((corec_bind i, NoSyn), (corec_def_bind i, corec_spec i T AT)))
49311
blanchet
parents: 49308
diff changeset
  1551
          ks Ts activeAs
blanchet
parents: 49308
diff changeset
  1552
      |>> apsnd split_list o split_list
blanchet
parents: 49308
diff changeset
  1553
      ||> `Local_Theory.restore;
blanchet
parents: 49308
diff changeset
  1554
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1555
    val phi = Proof_Context.export_morphism lthy_old lthy;
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
  1556
    val corecs = map (Morphism.term phi) corec_frees;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
  1557
    val corec_names = map (fst o dest_Const) corecs;
58443
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1558
    fun mk_corecs Ts passives actives =
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1559
      let val Tactives = map2 (curry mk_sumT) Ts actives;
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1560
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1561
        @{map 3} (fn name => fn T => fn active =>
58443
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1562
          Const (name, Library.foldr (op -->)
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1563
            (map2 (curry op -->) actives (mk_FTs (passives @ Tactives)), active --> T)))
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1564
        corec_names Ts actives
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  1565
      end;
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
  1566
    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
  1567
      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
55204
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1568
    val corec_defs = map (fn def =>
345ee77213b5 use Local_Theory.define instead of Specification.definition for internal constants
traytel
parents: 55197
diff changeset
  1569
      mk_unabs_def n (Morphism.thm phi def RS meta_eq_to_obj_eq)) corec_def_frees;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1570
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1571
    val case_sums =
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1572
      map2 (fn T => fn i => mk_case_sum (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
  1573
    val dtor_corec_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1574
      let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1575
        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
  1576
          let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1577
            val lhs = dtor $ (mk_corec corec_ss i $ z);
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1578
            val rhs = Term.list_comb (corec_map, passive_ids @ case_sums) $ (corec_s $ z);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1579
          in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1580
            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
  1581
          end;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1582
        val goals = @{map 5} 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
  1583
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1584
        @{map 3} (fn goal => fn unfold => fn map_cong0 =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  1585
          Goal.prove_sorry lthy [] [] goal
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1586
            (fn {context = ctxt, prems = _} => mk_corec_tac ctxt m corec_defs unfold map_cong0
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1587
              corec_Inl_sum_thms)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1588
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1589
          |> singleton (Proof_Context.export names_lthy lthy))
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1590
        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
  1591
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1592
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1593
    val corec_unique_mor_thm =
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1594
      let
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1595
        val id_fs = map2 (fn T => fn f => mk_case_sum (HOLogic.id_const T, f)) Ts unfold_fs;
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1596
        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
  1597
        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
  1598
        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
  1599
          (map2 mk_fun_eq unfold_fs ks));
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1600
      in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1601
        Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, unique))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1602
          (fn {context = ctxt, prems = _} => mk_corec_unique_mor_tac ctxt corec_defs
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  1603
            corec_Inl_sum_thms unfold_unique_mor_thm)
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1604
          |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1605
          |> singleton (Proof_Context.export names_lthy lthy)
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1606
      end;
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1607
53696
ee9eaab634e5 don't unfold as eager as in 11a77e4aa98b
traytel
parents: 53588
diff changeset
  1608
    val map_id0s_o_id =
ee9eaab634e5 don't unfold as eager as in 11a77e4aa98b
traytel
parents: 53588
diff changeset
  1609
      map (fn thm =>
55067
a452de24a877 tuned names
blanchet
parents: 55062
diff changeset
  1610
        mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "op o", OF _ refl]}) @{thm id_comp})
53696
ee9eaab634e5 don't unfold as eager as in 11a77e4aa98b
traytel
parents: 53588
diff changeset
  1611
      map_id0s;
ee9eaab634e5 don't unfold as eager as in 11a77e4aa98b
traytel
parents: 53588
diff changeset
  1612
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52912
diff changeset
  1613
    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
  1614
      `split_conj_thm (split_conj_prems n
52904
traytel
parents: 52839
diff changeset
  1615
        (mor_UNIV_thm RS iffD2 RS corec_unique_mor_thm)
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57726
diff changeset
  1616
        |> unfold_thms lthy (@{thms o_case_sum comp_id id_comp comp_assoc[symmetric]
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1617
           case_sum_o_inj(1)} @ map_id0s_o_id @ sym_map_comps)
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55413
diff changeset
  1618
        OF replicate n @{thm arg_cong2[of _ _ _ _ case_sum, OF refl]});
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 51551
diff changeset
  1619
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1620
    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
  1621
55644
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  1622
    val (coinduct_params, dtor_coinduct_thm) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1623
      let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1624
        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
  1625
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1626
        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
  1627
        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1628
          (@{map 3} mk_concl phis Jzs1 Jzs2));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1629
53105
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 53104
diff changeset
  1630
        fun mk_rel_prem 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
  1631
          let
55541
fd9ea8ae28f6 syntactic simplifications of internal (co)datatype constructions
traytel
parents: 55538
diff changeset
  1632
            val concl = Term.list_comb (rel, passive_eqs @ phis) $
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1633
              (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
  1634
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1635
            HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1636
              (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
  1637
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1638
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1639
        val rel_prems = @{map 5} mk_rel_prem phis dtors rels Jzs Jzs_copy;
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1640
        val dtor_coinduct_goal = Logic.list_implies (rel_prems, concl);
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1641
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1642
        val dtor_coinduct =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  1643
          Goal.prove_sorry lthy [] [] dtor_coinduct_goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1644
            (fn {context = ctxt, prems = _} => mk_dtor_coinduct_tac ctxt m raw_coind_thm bis_rel_thm
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1645
              rel_congs)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1646
          |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1647
          |> singleton (Proof_Context.export names_lthy lthy);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1648
      in
55644
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  1649
        (rev (Term.add_tfrees dtor_coinduct_goal []), dtor_coinduct)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1650
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1651
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1652
    val timer = time (timer "coinduction");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1653
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53258
diff changeset
  1654
    fun mk_dtor_map_DEADID_thm dtor_inject map_id0 =
c8628119d18e renamed BNF axiom
blanchet
parents: 53258
diff changeset
  1655
      trans OF [iffD2 OF [dtor_inject, id_apply], map_id0 RS sym];
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1656
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1657
    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
  1658
      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
  1659
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1660
    val JphiTs = map2 mk_pred2T passiveAs passiveBs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1661
    val Jpsi1Ts = map2 mk_pred2T passiveAs passiveCs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1662
    val Jpsi2Ts = map2 mk_pred2T passiveCs passiveBs;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1663
    val prodTsTs' = map2 (curry HOLogic.mk_prodT) Ts Ts';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1664
    val fstsTsTs' = map fst_const prodTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1665
    val sndsTsTs' = map snd_const prodTsTs';
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1666
    val activephiTs = map2 mk_pred2T activeAs activeBs;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1667
    val activeJphiTs = map2 mk_pred2T Ts Ts';
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1668
    val (((((Jphis, Jpsi1s), Jpsi2s), activephis), activeJphis), names_lthy) = names_lthy
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1669
      |> mk_Frees "R" JphiTs
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1670
      ||>> mk_Frees "R" Jpsi1Ts
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1671
      ||>> mk_Frees "Q" Jpsi2Ts
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  1672
      ||>> mk_Frees "S" activephiTs
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1673
      ||>> mk_Frees "JR" activeJphiTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  1674
    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
  1675
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1676
    fun mk_Jrel_DEADID_coinduct_thm () =
58579
b7bc5ba2f3fb rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
desharna
parents: 58578
diff changeset
  1677
      mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis (map HOLogic.eq_const Ts) Jphis
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1678
        Jzs Jz's dtors dtor's (fn {context = ctxt, prems} =>
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1679
          (unfold_thms_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1680
          REPEAT_DETERM (rtac ctxt allI 1) THEN rtac ctxt (dtor_coinduct_thm OF prems) 1)) lthy;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1681
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1682
    (*register new codatatypes as BNFs*)
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  1683
    val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jmap_unique_thms, dtor_Jset_thmss',
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57631
diff changeset
  1684
      dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, dtor_Jset_induct_thms, lthy) =
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  1685
      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
  1686
        (timer, replicate n DEADID_bnf,
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  1687
        map_split (`(mk_pointfree lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_ids), [],
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1688
        replicate n [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs,
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57631
diff changeset
  1689
        mk_Jrel_DEADID_coinduct_thm (), [], [], lthy)
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49584
diff changeset
  1690
      else let
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1691
        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
  1692
        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
  1693
        val uTs = map2 (curry op -->) Ts Ts';
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1694
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1695
        val (((((((((fs, fs'), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)),
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1696
          (ys_copy, ys'_copy)), Kss), names_lthy) = names_lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1697
          |> 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
  1698
          ||>> 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
  1699
          ||>> mk_Frees "g" gTs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1700
          ||>> mk_Frees "u" uTs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1701
          ||>> mk_Frees' "b" Ts'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1702
          ||>> mk_Frees' "b" Ts'
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1703
          ||>> mk_Frees' "y" passiveAs
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1704
          ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) Ts) ATs);;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1705
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1706
        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
  1707
          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
  1708
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1709
        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
  1710
          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
  1711
        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
  1712
        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
  1713
          mk_unfold Ts' (map2 (fn dtor => fn Fmap =>
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1714
            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
  1715
        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
  1716
        val mk_mapsAB = mk_maps passiveAs passiveBs;
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49499
diff changeset
  1717
        val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1718
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1719
        val set_bss =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1720
          map (flat o map2 (fn B => fn b =>
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1721
            if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1722
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1723
        fun col_bind j = mk_internal_b (colN ^ (if m = 1 then "" else string_of_int j));
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1724
        val col_def_bind = rpair [] o Thm.def_binding o col_bind;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1725
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1726
        fun col_spec j Zero hrec hrec' =
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1727
          let
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1728
            fun mk_Suc dtor sets z z' =
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1729
              let
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1730
                val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m sets);
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1731
                fun mk_UN set k = mk_UNION (set $ (dtor $ z)) (mk_nthN n hrec k);
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1732
              in
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1733
                Term.absfree z'
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1734
                  (mk_union (set $ (dtor $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1735
              end;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1736
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1737
            val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1738
              (HOLogic.mk_tuple (@{map 4} mk_Suc dtors FTs_setss Jzs Jzs')));
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1739
          in
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1740
            mk_rec_nat Zero Suc
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1741
          end;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1742
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1743
        val ((col_frees, (_, col_def_frees)), (lthy, lthy_old)) =
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1744
          lthy
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1745
          |> @{fold_map 4} (fn j => fn Zero => fn hrec => fn hrec' => Local_Theory.define
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1746
            ((col_bind j, NoSyn), (col_def_bind j, col_spec j Zero hrec hrec')))
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1747
            ls Zeros hrecs hrecs'
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1748
          |>> apsnd split_list o split_list
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1749
          ||> `Local_Theory.restore;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1750
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1751
        val phi = Proof_Context.export_morphism lthy_old lthy;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1752
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1753
        val col_defs = map (fn def => Morphism.thm phi def RS meta_eq_to_obj_eq) col_def_frees;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1754
        val cols = map (fst o Term.dest_Const o Morphism.term phi) col_frees;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1755
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1756
        fun mk_col Ts nat i j T =
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1757
          let
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1758
            val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) Ts)
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1759
            val colT = HOLogic.natT --> hrecT;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1760
          in
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1761
            mk_nthN n (Term.list_comb (Const (nth cols (j - 1), colT), [nat])) i
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1762
          end;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1763
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1764
        val col_0ss = mk_rec_simps n @{thm rec_nat_0_imp} col_defs;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1765
        val col_Sucss = mk_rec_simps n @{thm rec_nat_Suc_imp} col_defs;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1766
        val col_0ss' = transpose col_0ss;
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1767
        val col_Sucss' = transpose col_Sucss;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1768
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1769
        fun mk_set Ts i j T =
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1770
          Abs (Name.uu, nth Ts (i - 1), mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1771
            (Term.absfree nat' (mk_col Ts nat i j T $ Bound 1)));
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  1772
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1773
        val setss = map (fn i => map2 (mk_set Ts i) ls passiveAs) ks;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1774
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1775
        val (Jbnf_consts, lthy) =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1776
          @{fold_map 7} (fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets => fn T =>
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1777
              fn lthy =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1778
            define_bnf_consts Hardly_Inline (user_policy Note_Some lthy) false (SOME deads)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1779
              map_b rel_b set_bs
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1780
              ((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd),
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1781
                [Const (@{const_name undefined}, T)]), NONE) lthy)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1782
          bs map_bs rel_bs set_bss fs_maps setss Ts lthy;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1783
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1784
        val (_, Jconsts, Jconst_defs, mk_Jconsts) = @{split_list 4} Jbnf_consts;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1785
        val (_, Jsetss, Jbds_Ds, _, _) = @{split_list 5} Jconsts;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1786
        val (Jmap_defs, Jset_defss, Jbd_defs, _, Jrel_defs) = @{split_list 5} Jconst_defs;
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1787
        val (mk_Jmaps_Ds, mk_Jt_Ds, _, mk_Jrels_Ds, _) = @{split_list 5} mk_Jconsts;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1788
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1789
        val Jrel_unabs_defs = map (fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) Jrel_defs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1790
        val Jset_defs = flat Jset_defss;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1791
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1792
        fun mk_Jmaps As Bs = map (fn mk => mk deads As Bs) mk_Jmaps_Ds;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1793
        fun mk_Jsetss As = map2 (fn mk => fn Jsets => map (mk deads As) Jsets) mk_Jt_Ds Jsetss;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1794
        val Jbds = map2 (fn mk => mk deads passiveAs) mk_Jt_Ds Jbds_Ds;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1795
        fun mk_Jrels As Bs = map (fn mk => mk deads As Bs) mk_Jrels_Ds;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1796
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1797
        val Jmaps = mk_Jmaps passiveAs passiveBs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1798
        val fs_Jmaps = map (fn m => Term.list_comb (m, fs)) Jmaps;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1799
        val fs_copy_Jmaps = map (fn m => Term.list_comb (m, fs_copy)) Jmaps;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1800
        val gs_Jmaps = map (fn m => Term.list_comb (m, gs)) (mk_Jmaps passiveBs passiveCs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1801
        val fgs_Jmaps = map (fn m => Term.list_comb (m, map2 (curry HOLogic.mk_comp) gs fs))
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1802
          (mk_Jmaps passiveAs passiveCs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1803
        val (Jsetss_by_range, Jsetss_by_bnf) = `transpose (mk_Jsetss passiveAs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1804
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1805
        val timer = time (timer "bnf constants for the new datatypes");
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1806
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1807
        val (dtor_Jmap_thms, Jmap_thms) =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1808
          let
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1809
            fun mk_goal fs_Jmap map dtor dtor' = mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_Jmap),
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1810
              HOLogic.mk_comp (Term.list_comb (map, fs @ fs_Jmaps), dtor));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1811
            val goals = @{map 4} mk_goal fs_Jmaps map_FTFT's dtors dtor's;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1812
            val maps =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1813
              @{map 5} (fn goal => fn unfold => fn map_comp => fn map_cong0 => fn map_arg_cong =>
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1814
                Goal.prove_sorry lthy [] [] goal
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1815
                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1816
                     mk_map_tac ctxt m n map_arg_cong unfold map_comp map_cong0)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1817
                |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1818
                |> singleton (Proof_Context.export names_lthy lthy))
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  1819
              goals dtor_unfold_thms map_comps map_cong0s map_arg_cong_thms;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1820
          in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1821
            map_split (fn thm => (thm RS @{thm comp_eq_dest}, thm)) maps
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1822
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1823
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  1824
        val (dtor_Jmap_unique_thms, dtor_Jmap_unique_thm) =
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1825
          let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1826
            fun mk_prem u map dtor dtor' =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1827
              mk_Trueprop_eq (HOLogic.mk_comp (dtor', u),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1828
                HOLogic.mk_comp (Term.list_comb (map, fs @ us), dtor));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1829
            val prems = @{map 4} mk_prem us map_FTFT's dtors dtor's;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1830
            val goal =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1831
              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1832
                (map2 (curry HOLogic.mk_eq) us fs_Jmaps));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1833
          in
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  1834
            `split_conj_thm (Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, goal))
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1835
              (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1836
                mk_dtor_map_unique_tac ctxt dtor_unfold_unique_thm sym_map_comps)
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1837
            |> Thm.close_derivation
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  1838
            |> singleton (Proof_Context.export names_lthy lthy))
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1839
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1840
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1841
        val Jmap_comp0_thms =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1842
          let
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1843
            val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1844
              (@{map 3} (fn fmap => fn gmap => fn fgmap =>
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1845
                 HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap))
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1846
              fs_Jmaps gs_Jmaps fgs_Jmaps))
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1847
          in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1848
            split_conj_thm (Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1849
              (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1850
                mk_map_comp0_tac ctxt Jmap_thms map_comp0s dtor_Jmap_unique_thm)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1851
              |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1852
              |> singleton (Proof_Context.export names_lthy lthy))
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1853
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1854
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1855
        val timer = time (timer "map functions for the new codatatypes");
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1856
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1857
        val Jset_minimal_thms =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1858
          let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1859
            fun mk_passive_prem set dtor x K =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1860
              Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (K $ x)));
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1861
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1862
            fun mk_active_prem dtor x1 K1 set x2 K2 =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1863
              fold_rev Logic.all [x1, x2]
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1864
                (Logic.mk_implies (mk_Trueprop_mem (x2, set $ (dtor $ x1)),
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1865
                  HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1))));
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1866
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1867
            val premss = map2 (fn j => fn Ks =>
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1868
              @{map 4} mk_passive_prem (map (fn xs => nth xs (j - 1)) FTs_setss) dtors Jzs Ks @
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1869
                flat (@{map 4} (fn sets => fn s => fn x1 => fn K1 =>
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1870
                  @{map 3} (mk_active_prem s x1 K1) (drop m sets) Jzs_copy Ks) FTs_setss dtors Jzs Ks))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1871
              ls Kss;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1872
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1873
            val col_minimal_thms =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1874
              let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1875
                fun mk_conjunct j T i K x = mk_leq (mk_col Ts nat i j T $ x) (K $ x);
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1876
                fun mk_concl j T Ks = list_all_free Jzs
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1877
                  (Library.foldr1 HOLogic.mk_conj (@{map 3} (mk_conjunct j T) ks Ks Jzs));
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1878
                val concls = @{map 3} mk_concl ls passiveAs Kss;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1879
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1880
                val goals = map2 (fn prems => fn concl =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1881
                  Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1882
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1883
                val ctss =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1884
                  map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) concls;
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1885
              in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1886
                @{map 4} (fn goal => fn cts => fn col_0s => fn col_Sucs =>
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1887
                  Goal.prove_sorry lthy [] [] goal
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1888
                    (fn {context = ctxt, prems = _} => mk_col_minimal_tac ctxt m cts col_0s
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1889
                      col_Sucs)
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1890
                  |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1891
                  |> singleton (Proof_Context.export names_lthy lthy))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1892
                goals ctss col_0ss' col_Sucss'
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1893
              end;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1894
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1895
            fun mk_conjunct set K x = mk_leq (set $ x) (K $ x);
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1896
            fun mk_concl sets Ks = Library.foldr1 HOLogic.mk_conj (@{map 3} mk_conjunct sets Ks Jzs);
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1897
            val concls = map2 mk_concl Jsetss_by_range Kss;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1898
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1899
            val goals = map2 (fn prems => fn concl =>
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1900
              Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls;
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1901
          in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1902
            map2 (fn goal => fn col_minimal =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1903
              Goal.prove_sorry lthy [] [] goal
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1904
                (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1905
                  mk_Jset_minimal_tac ctxt n col_minimal)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1906
              |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1907
              |> singleton (Proof_Context.export names_lthy lthy))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1908
            goals col_minimal_thms
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1909
          end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1910
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1911
        val (dtor_Jset_incl_thmss, dtor_set_Jset_incl_thmsss) =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1912
          let
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1913
            fun mk_set_incl_Jset dtor x set Jset =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1914
              HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (Jset $ x));
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1915
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1916
            fun mk_set_Jset_incl_Jset dtor x y set Jset1 Jset2 =
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1917
              Logic.mk_implies (mk_Trueprop_mem (x, set $ (dtor $ y)),
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1918
                HOLogic.mk_Trueprop (mk_leq (Jset1 $ x) (Jset2 $ y)));
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1919
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1920
            val set_incl_Jset_goalss =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1921
              @{map 4} (fn dtor => fn x => fn sets => fn Jsets =>
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1922
                map2 (mk_set_incl_Jset dtor x) (take m sets) Jsets)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1923
              dtors Jzs FTs_setss Jsetss_by_bnf;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1924
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1925
            (*x(k) : F(i)set(m+k) (dtor(i) y(i)) ==> J(k)set(j) x(k) <= J(i)set(j) y(i)*)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1926
            val set_Jset_incl_Jset_goalsss =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1927
              @{map 4} (fn dtori => fn yi => fn sets => fn Jsetsi =>
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1928
                @{map 3} (fn xk => fn set => fn Jsetsk =>
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1929
                  map2 (mk_set_Jset_incl_Jset dtori xk yi set) Jsetsk Jsetsi)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1930
                Jzs_copy (drop m sets) Jsetss_by_bnf)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1931
              dtors Jzs FTs_setss Jsetss_by_bnf;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1932
          in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1933
            (map2 (fn goals => fn rec_Sucs =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1934
              map2 (fn goal => fn rec_Suc =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1935
                Goal.prove_sorry lthy [] [] goal
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1936
                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1937
                    mk_set_incl_Jset_tac ctxt rec_Suc)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1938
                |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1939
                |> singleton (Proof_Context.export names_lthy lthy))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1940
              goals rec_Sucs)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1941
            set_incl_Jset_goalss col_Sucss,
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1942
            map2 (fn goalss => fn rec_Sucs =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1943
              map2 (fn k => fn goals =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1944
                map2 (fn goal => fn rec_Suc =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1945
                  Goal.prove_sorry lthy [] [] goal
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1946
                    (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1947
                      mk_set_Jset_incl_Jset_tac ctxt n rec_Suc k)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1948
                  |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  1949
                  |> singleton (Proof_Context.export names_lthy lthy))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1950
                goals rec_Sucs)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1951
              ks goalss)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1952
            set_Jset_incl_Jset_goalsss col_Sucss)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1953
          end;
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1954
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1955
        val set_incl_Jset_thmss' = transpose dtor_Jset_incl_thmss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1956
        val set_Jset_incl_Jset_thmsss' = transpose (map transpose dtor_set_Jset_incl_thmsss);
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1957
        val set_Jset_thmss = map (map (fn thm => thm RS @{thm set_mp})) dtor_Jset_incl_thmss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1958
        val set_Jset_Jset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1959
          dtor_set_Jset_incl_thmsss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1960
        val set_Jset_thmss' = transpose set_Jset_thmss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1961
        val set_Jset_Jset_thmsss' = transpose (map transpose set_Jset_Jset_thmsss);
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  1962
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1963
        val dtor_Jset_induct_thms =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1964
          let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1965
            val incls =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1966
              maps (map (fn thm => thm RS @{thm subset_Collect_iff})) dtor_Jset_incl_thmss @
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1967
                @{thms subset_Collect_iff[OF subset_refl]};
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1968
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1969
            val cTs = map (SOME o Thm.ctyp_of lthy) params';
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1970
            fun mk_induct_tinst phis jsets y y' =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1971
              @{map 4} (fn phi => fn jset => fn Jz => fn Jz' =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1972
                SOME (Thm.cterm_of lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y',
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1973
                  HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz))))))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1974
              phis jsets Jzs Jzs';
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1975
          in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1976
            @{map 6} (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1977
              ((set_minimal
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1978
                |> Drule.instantiate' cTs (mk_induct_tinst phis jsets y y')
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1979
                |> unfold_thms lthy incls) OF
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1980
                (replicate n ballI @
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1981
                  maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1982
              |> singleton (Proof_Context.export names_lthy lthy)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  1983
              |> rule_by_tactic lthy (ALLGOALS (TRY o etac lthy asm_rl)))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1984
            Jset_minimal_thms set_Jset_incl_Jset_thmsss' Jsetss_by_range ys ys' dtor_set_induct_phiss
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1985
          end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  1986
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1987
        val (dtor_Jset_thmss', dtor_Jset_thmss) =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1988
          let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1989
            fun mk_simp_goal relate pas_set act_sets sets dtor z set =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1990
              relate (set $ z, mk_union (pas_set $ (dtor $ z),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1991
                 Library.foldl1 mk_union
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1992
                   (map2 (fn X => mk_UNION (X $ (dtor $ z))) act_sets sets)));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1993
            fun mk_goals eq =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1994
              map2 (fn i => fn sets =>
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  1995
                @{map 4} (fn Fsets =>
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1996
                  mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets)
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1997
                FTs_setss dtors Jzs sets)
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1998
              ls Jsetss_by_range;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  1999
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2000
            val le_goals = map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2001
              (mk_goals (uncurry mk_leq));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2002
            val set_le_thmss = map split_conj_thm
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2003
              (@{map 4} (fn goal => fn Jset_minimal => fn set_Jsets => fn set_Jset_Jsetss =>
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2004
                Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2005
                  (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2006
                    mk_set_le_tac ctxt n Jset_minimal set_Jsets set_Jset_Jsetss)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2007
                |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2008
                |> singleton (Proof_Context.export names_lthy lthy))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2009
              le_goals Jset_minimal_thms set_Jset_thmss' set_Jset_Jset_thmsss');
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2010
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2011
            val ge_goalss = map (map HOLogic.mk_Trueprop) (mk_goals (uncurry mk_leq o swap));
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  2012
            val set_ge_thmss =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2013
              @{map 3} (@{map 3} (fn goal => fn set_incl_Jset => fn set_Jset_incl_Jsets =>
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2014
                Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2015
                  (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2016
                    mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2017
                |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2018
                |> singleton (Proof_Context.export names_lthy lthy)))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2019
              ge_goalss set_incl_Jset_thmss' set_Jset_incl_Jset_thmsss'
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2020
          in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2021
            map2 (map2 (fn le => fn ge => equalityI OF [le, ge])) set_le_thmss set_ge_thmss
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2022
            |> `transpose
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2023
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2024
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2025
        val timer = time (timer "set functions for the new codatatypes");
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2026
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2027
        val colss = map2 (fn j => fn T =>
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  2028
          map (fn i => mk_col Ts nat i j T) ks) ls passiveAs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2029
        val colss' = map2 (fn j => fn T =>
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  2030
          map (fn i => mk_col Ts' nat i j T) ks) ls passiveBs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2031
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2032
        val col_natural_thmss =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2033
          let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2034
            fun mk_col_natural f map z col col' =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2035
              HOLogic.mk_eq (mk_image f $ (col $ z), col' $ (map $ z));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2036
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2037
            fun mk_goal f cols cols' = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2038
              (@{map 4} (mk_col_natural f) fs_Jmaps Jzs cols cols'));
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2039
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2040
            val goals = @{map 3} mk_goal fs colss colss';
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2041
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2042
            val ctss =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2043
              map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) goals;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2044
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2045
            val thms =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2046
              @{map 4} (fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2047
                Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2048
                  (fn {context = ctxt, prems = _} => mk_col_natural_tac ctxt cts rec_0s rec_Sucs
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2049
                    dtor_Jmap_thms set_mapss)
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2050
                |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2051
                |> singleton (Proof_Context.export names_lthy lthy))
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  2052
              goals ctss col_0ss' col_Sucss';
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2053
          in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2054
            map (split_conj_thm o mk_specN n) thms
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2055
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2056
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2057
        val col_bd_thmss =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2058
          let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2059
            fun mk_col_bd z col bd = mk_ordLeq (mk_card_of (col $ z)) bd;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2060
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2061
            fun mk_goal bds cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2062
              (@{map 3} mk_col_bd Jzs cols bds));
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2063
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2064
            val goals = map (mk_goal Jbds) colss;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2065
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2066
            val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat])
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2067
              (map (mk_goal (replicate n sbd)) colss);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2068
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2069
            val thms =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2070
              @{map 5} (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2071
                Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2072
                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2073
                    mk_col_bd_tac ctxt m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2074
                |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2075
                |> singleton (Proof_Context.export names_lthy lthy))
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  2076
              ls goals ctss col_0ss' col_Sucss';
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2077
          in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2078
            map (split_conj_thm o mk_specN n) thms
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2079
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2080
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2081
        val map_cong0_thms =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2082
          let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2083
            val cTs = map (SOME o Thm.ctyp_of lthy o
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2084
              Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2085
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2086
            fun mk_prem z set f g y y' =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2087
              mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2088
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2089
            fun mk_prems sets z =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2090
              Library.foldr1 HOLogic.mk_conj (@{map 5} (mk_prem z) sets fs fs_copy ys ys')
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2091
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2092
            fun mk_map_cong0 sets z fmap gmap =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2093
              HOLogic.mk_imp (mk_prems sets z, HOLogic.mk_eq (fmap $ z, gmap $ z));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2094
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2095
            fun mk_coind_body sets (x, T) z fmap gmap y y_copy =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2096
              HOLogic.mk_conj
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2097
                (HOLogic.mk_mem (z, HOLogic.mk_Collect (x, T, mk_prems sets z)),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2098
                  HOLogic.mk_conj (HOLogic.mk_eq (y, fmap $ z),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2099
                    HOLogic.mk_eq (y_copy, gmap $ z)))
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2100
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2101
            fun mk_cphi sets (z' as (x, T)) z fmap gmap y' y y'_copy y_copy =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2102
              HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy)
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2103
              |> Term.absfree y'_copy
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2104
              |> Term.absfree y'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2105
              |> Thm.cterm_of lthy;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2106
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2107
            val cphis = @{map 9} mk_cphi
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2108
              Jsetss_by_bnf Jzs' Jzs fs_Jmaps fs_copy_Jmaps Jys' Jys Jys'_copy Jys_copy;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2109
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2110
            val coinduct = Drule.instantiate' cTs (map SOME cphis) dtor_coinduct_thm;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2111
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2112
            val goal =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2113
              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2114
                (@{map 4} mk_map_cong0 Jsetss_by_bnf Jzs fs_Jmaps fs_copy_Jmaps));
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2115
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2116
            val thm =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2117
              Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2118
                (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt m (rtac ctxt coinduct) map_comps
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2119
                  dtor_Jmap_thms map_cong0s
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2120
                  set_mapss set_Jset_thmss set_Jset_Jset_thmsss in_rels)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2121
              |> Thm.close_derivation
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2122
              |>  singleton (Proof_Context.export names_lthy lthy);
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2123
          in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2124
            split_conj_thm thm
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2125
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2126
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2127
        val in_Jrels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2128
            Jrel_unabs_defs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2129
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2130
        val Jrels = mk_Jrels passiveAs passiveBs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2131
        val Jrelphis = map (fn rel => Term.list_comb (rel, Jphis)) Jrels;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2132
        val relphis = map (fn rel => Term.list_comb (rel, Jphis @ Jrelphis)) rels;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2133
        val Jrelpsi1s = map (fn rel => Term.list_comb (rel, Jpsi1s)) (mk_Jrels passiveAs passiveCs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2134
        val Jrelpsi2s = map (fn rel => Term.list_comb (rel, Jpsi2s)) (mk_Jrels passiveCs passiveBs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2135
        val Jrelpsi12s = map (fn rel =>
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2136
            Term.list_comb (rel, map2 (curry mk_rel_compp) Jpsi1s Jpsi2s)) Jrels;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2137
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51869
diff changeset
  2138
        val dtor_Jrel_thms =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2139
          let
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2140
            fun mk_goal Jz Jz' dtor dtor' Jrelphi relphi =
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2141
              mk_Trueprop_eq (Jrelphi $ Jz $ Jz', relphi $ (dtor $ Jz) $ (dtor' $ Jz'));
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2142
            val goals = @{map 6} 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
  2143
          in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2144
            @{map 12} (fn i => fn goal => fn in_rel => fn map_comp0 => 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
  2145
              fn dtor_map => fn dtor_sets => fn dtor_inject => fn dtor_ctor =>
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  2146
              fn set_map0s => fn dtor_set_incls => fn dtor_set_set_inclss =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51447
diff changeset
  2147
              Goal.prove_sorry lthy [] [] goal
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
  2148
                (K (mk_dtor_rel_tac lthy in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  2149
                  dtor_inject dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss))
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2150
              |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2151
              |> singleton (Proof_Context.export names_lthy lthy))
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2152
            ks goals in_rels map_comps map_cong0s dtor_Jmap_thms dtor_Jset_thmss'
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2153
              dtor_inject_thms dtor_ctor_thms set_mapss dtor_Jset_incl_thmss
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2154
              dtor_set_Jset_incl_thmsss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2155
          end;
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 passiveABs = map2 (curry HOLogic.mk_prodT) passiveAs passiveBs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2158
      val zip_ranTs = passiveABs @ prodTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2159
      val allJphis = Jphis @ activeJphis;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2160
      val zipFTs = mk_FTs zip_ranTs;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2161
      val zipTs = @{map 3} (fn T => fn T' => fn FT => T --> T' --> FT) Ts Ts' zipFTs;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2162
      val zip_zTs = mk_Ts passiveABs;
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2163
      val (((zips, (abs, abs')), (zip_zs, zip_zs')), names_lthy) = names_lthy
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2164
        |> mk_Frees "zip" zipTs
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2165
        ||>> mk_Frees' "ab" passiveABs
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2166
        ||>> mk_Frees' "z" zip_zTs;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2167
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2168
      val Iphi_sets =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2169
        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
  2170
      val in_phis = map2 (mk_in Iphi_sets) (mk_setss zip_ranTs) zipFTs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2171
      val fstABs = map fst_const passiveABs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2172
      val all_fsts = fstABs @ fstsTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2173
      val map_all_fsts = map2 (fn Ds => fn bnf =>
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2174
        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveAs @ Ts) bnf, all_fsts)) Dss bnfs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2175
      val Jmap_fsts = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2176
        else Term.list_comb (map, fstABs)) (mk_Jmaps passiveABs passiveAs) Ts;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2177
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2178
      val sndABs = map snd_const passiveABs;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2179
      val all_snds = sndABs @ sndsTsTs';
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2180
      val map_all_snds = map2 (fn Ds => fn bnf =>
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2181
        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveBs @ Ts') bnf, all_snds)) Dss bnfs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2182
      val Jmap_snds = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2183
        else Term.list_comb (map, sndABs)) (mk_Jmaps passiveABs passiveBs) Ts;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2184
      val zip_unfolds = map (mk_unfold zip_zTs (map HOLogic.mk_split zips)) ks;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2185
      val zip_setss = mk_Jsetss passiveABs |> transpose;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2186
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  2187
      fun Jrel_coinduct_tac {context = ctxt, prems = CIHs} =
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2188
        let
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2189
          fun mk_helper_prem phi in_phi zip x y map map' dtor dtor' =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2190
            let
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2191
              val zipxy = zip $ x $ y;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2192
            in
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2193
              HOLogic.mk_Trueprop (list_all_free [x, y]
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2194
                (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
  2195
                  HOLogic.mk_conj (HOLogic.mk_eq (map $ zipxy, dtor $ x),
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2196
                    HOLogic.mk_eq (map' $ zipxy, dtor' $ y))))))
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2197
            end;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2198
          val helper_prems = @{map 9} mk_helper_prem
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2199
            activeJphis in_phis zips Jzs Jz's map_all_fsts map_all_snds dtors dtor's;
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2200
          fun mk_helper_coind_phi fst phi x alt y map zip_unfold =
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2201
            list_exists_free [if fst then y else x] (HOLogic.mk_conj (phi $ x $ y,
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2202
              HOLogic.mk_eq (alt, map $ (zip_unfold $ HOLogic.mk_prod (x, y)))))
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2203
          val coind1_phis = @{map 6} (mk_helper_coind_phi true)
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2204
            activeJphis Jzs Jzs_copy Jz's Jmap_fsts zip_unfolds;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2205
          val coind2_phis = @{map 6} (mk_helper_coind_phi false)
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2206
              activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds;
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2207
          fun mk_cts zs z's phis =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2208
            @{map 3} (fn z => fn z' => fn phi =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2209
              SOME (Thm.cterm_of lthy (fold_rev (Term.absfree o Term.dest_Free) [z', z] phi)))
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2210
            zs z's phis @
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2211
            map (SOME o Thm.cterm_of lthy) (splice z's zs);
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2212
          val cts1 = mk_cts Jzs Jzs_copy coind1_phis;
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2213
          val cts2 = mk_cts Jz's Jz's_copy coind2_phis;
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2214
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2215
          fun mk_helper_coind_concl z alt coind_phi =
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2216
            HOLogic.mk_imp (coind_phi, HOLogic.mk_eq (alt, z));
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2217
          val helper_coind1_concl =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2218
            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2219
              (@{map 3} mk_helper_coind_concl Jzs Jzs_copy coind1_phis));
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2220
          val helper_coind2_concl =
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2221
            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2222
              (@{map 3} mk_helper_coind_concl Jz's Jz's_copy coind2_phis));
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2223
55644
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  2224
          fun mk_helper_coind_thms fst concl cts =
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2225
            Goal.prove_sorry lthy [] [] (Logic.list_implies (helper_prems, concl))
55644
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  2226
              (fn {context = ctxt, prems = _} => mk_rel_coinduct_coind_tac ctxt fst m
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  2227
                (cterm_instantiate_pos cts dtor_coinduct_thm) ks map_comps map_cong0s
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  2228
                map_arg_cong_thms set_mapss dtor_unfold_thms dtor_Jmap_thms in_rels)
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2229
            |> Thm.close_derivation
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2230
            |> split_conj_thm
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2231
            |> Proof_Context.export names_lthy lthy;
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2232
55644
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  2233
          val helper_coind1_thms = mk_helper_coind_thms true helper_coind1_concl cts1;
b657146dc030 only one internal coinduction rule
traytel
parents: 55642
diff changeset
  2234
          val helper_coind2_thms = mk_helper_coind_thms false helper_coind2_concl cts2;
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2235
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2236
          fun mk_helper_ind_phi phi ab fst snd z active_phi x y zip_unfold =
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2237
            list_all_free [x, y] (HOLogic.mk_imp
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2238
              (HOLogic.mk_conj (active_phi $ x $ y,
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2239
                 HOLogic.mk_eq (z, zip_unfold $ HOLogic.mk_prod (x, y))),
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2240
              phi $ (fst $ ab) $ (snd $ ab)));
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2241
          val helper_ind_phiss =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2242
            @{map 4} (fn Jphi => fn ab => fn fst => fn snd =>
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2243
              @{map 5} (mk_helper_ind_phi Jphi ab fst snd)
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2244
              zip_zs activeJphis Jzs Jz's zip_unfolds)
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2245
            Jphis abs fstABs sndABs;
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2246
          val ctss = map2 (fn ab' => fn phis =>
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2247
              map2 (fn z' => fn phi =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2248
                SOME (Thm.cterm_of lthy (Term.absfree ab' (Term.absfree z' phi))))
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2249
              zip_zs' phis @
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2250
              map (SOME o Thm.cterm_of lthy) zip_zs)
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2251
            abs' helper_ind_phiss;
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2252
          fun mk_helper_ind_concl ab' z ind_phi set =
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2253
            mk_Ball (set $ z) (Term.absfree ab' ind_phi);
57567
d554b0097ad4 add mk_Trueprop_mem utility function
desharna
parents: 57093
diff changeset
  2254
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2255
          val mk_helper_ind_concls =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2256
            @{map 3} (fn ab' => fn ind_phis => fn zip_sets =>
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2257
              @{map 3} (mk_helper_ind_concl ab') zip_zs ind_phis zip_sets)
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2258
            abs' helper_ind_phiss zip_setss
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2259
            |> map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj);
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2260
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2261
          val helper_ind_thmss = if m = 0 then replicate n [] else
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2262
            @{map 4} (fn concl => fn j => fn set_induct => fn cts =>
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2263
              Goal.prove_sorry lthy [] [] (Logic.list_implies (helper_prems, concl))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  2264
                (fn {context = ctxt, prems = _} => mk_rel_coinduct_ind_tac ctxt m ks
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2265
                  dtor_unfold_thms set_mapss j (cterm_instantiate_pos cts set_induct))
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2266
              |> Thm.close_derivation
55602
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2267
              |> split_conj_thm
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2268
              |> Proof_Context.export names_lthy lthy)
257bd115fcca made tactics more robust
traytel
parents: 55581
diff changeset
  2269
            mk_helper_ind_concls ls dtor_Jset_induct_thms ctss
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2270
            |> transpose;
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2271
        in
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  2272
          mk_rel_coinduct_tac ctxt CIHs in_rels in_Jrels
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2273
            helper_ind_thmss helper_coind1_thms helper_coind2_thms
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2274
        end;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2275
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2276
      val Jrel_coinduct_thm =
58579
b7bc5ba2f3fb rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
desharna
parents: 58578
diff changeset
  2277
        mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis Jrels Jphis Jzs Jz's dtors dtor's
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52344
diff changeset
  2278
          Jrel_coinduct_tac lthy;
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2279
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2280
        val le_Jrel_OO_thm =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2281
          let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2282
            fun mk_le_Jrel_OO Jrelpsi1 Jrelpsi2 Jrelpsi12 =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2283
              mk_leq (mk_rel_compp (Jrelpsi1, Jrelpsi2)) Jrelpsi12;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2284
            val goals = @{map 3} mk_le_Jrel_OO Jrelpsi1s Jrelpsi2s Jrelpsi12s;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2285
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2286
            val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2287
          in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2288
            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2289
              mk_le_rel_OO_tac ctxt Jrel_coinduct_thm dtor_Jrel_thms le_rel_OOs)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2290
            |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2291
            |> singleton (Proof_Context.export names_lthy lthy)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2292
          end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2293
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2294
        val timer = time (timer "helpers for BNF properties");
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2295
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2296
        fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit);
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2297
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2298
        val all_unitTs = replicate live HOLogic.unitT;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2299
        val unitTs = replicate n HOLogic.unitT;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2300
        val unit_funs = replicate n (Term.absdummy HOLogic.unitT HOLogic.unit);
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2301
        fun mk_map_args I =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2302
          map (fn i =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2303
            if member (op =) I i then Term.absdummy HOLogic.unitT (nth ys i)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2304
            else mk_undefined (HOLogic.unitT --> nth passiveAs i))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2305
          (0 upto (m - 1));
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2306
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2307
        fun mk_nat_wit Ds bnf (I, wit) () =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2308
          let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2309
            val passiveI = filter (fn i => i < m) I;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2310
            val map_args = mk_map_args passiveI;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2311
          in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2312
            Term.absdummy HOLogic.unitT (Term.list_comb
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2313
              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ wit)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2314
          end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2315
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2316
        fun mk_dummy_wit Ds bnf I =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2317
          let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2318
            val map_args = mk_map_args I;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2319
          in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2320
            Term.absdummy HOLogic.unitT (Term.list_comb
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2321
              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2322
              mk_undefined (mk_T_of_bnf Ds all_unitTs bnf))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2323
          end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2324
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2325
        val nat_witss =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2326
          map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2327
            (replicate (nwits_of_bnf bnf) (replicate live HOLogic.unitT)) bnf
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2328
            |> map (fn (I, wit) =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2329
              (I, Lazy.lazy (mk_nat_wit Ds bnf (I, Term.list_comb (wit, map (K HOLogic.unit) I))))))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2330
          Dss bnfs;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2331
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2332
        val nat_wit_thmss = map2 (curry op ~~) nat_witss (map wit_thmss_of_bnf bnfs)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2333
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2334
        val Iss = map (map fst) nat_witss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2335
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2336
        fun filter_wits (I, wit) =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2337
          let val J = filter (fn i => i < m) I;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2338
          in (J, (length J < length I, wit)) end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2339
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2340
        val wit_treess = map_index (fn (i, Is) =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2341
          map_index (finish Iss m [i+m] (i+m)) Is) Iss
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2342
          |> map (minimize_wits o map filter_wits o minimize_wits o flat);
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2343
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2344
        val coind_wit_argsss =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2345
          map (map (tree_to_coind_wits nat_wit_thmss o snd o snd) o filter (fst o snd)) wit_treess;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2346
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2347
        val nonredundant_coind_wit_argsss =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2348
          fold (fn i => fn argsss =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2349
            nth_map (i - 1) (filter_out (fn xs =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2350
              exists (fn ys =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2351
                let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2352
                  val xs' = (map (fst o fst) xs, snd (fst (hd xs)));
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2353
                  val ys' = (map (fst o fst) ys, snd (fst (hd ys)));
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2354
                in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2355
                  eq_pair (subset (op =)) (eq_set (op =)) (xs', ys') andalso not (fst xs' = fst ys')
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2356
                end)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2357
              (flat argsss)))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2358
            argsss)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2359
          ks coind_wit_argsss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2360
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2361
        fun prepare_args args =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2362
          let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2363
            val I = snd (fst (hd args));
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2364
            val (dummys, args') =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2365
              map_split (fn i =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2366
                (case find_first (fn arg => fst (fst arg) = i - 1) args of
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2367
                  SOME (_, ((_, wit), thms)) => (NONE, (Lazy.force wit, thms))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2368
                | NONE =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2369
                  (SOME (i - 1), (mk_dummy_wit (nth Dss (i - 1)) (nth bnfs (i - 1)) I, []))))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2370
              ks;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2371
          in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2372
            ((I, dummys), apsnd flat (split_list args'))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2373
          end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2374
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2375
        fun mk_coind_wits ((I, dummys), (args, thms)) =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2376
          ((I, dummys), (map (fn i => mk_unfold Ts args i $ HOLogic.unit) ks, thms));
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2377
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2378
        val coind_witss =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2379
          maps (map (mk_coind_wits o prepare_args)) nonredundant_coind_wit_argsss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2380
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2381
        val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2382
          (replicate (nwits_of_bnf bnf) Ds)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2383
          (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2384
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2385
        val ctor_witss =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2386
          map (map (uncurry close_wit o tree_to_ctor_wit ys ctors witss o snd o snd) o
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2387
            filter_out (fst o snd)) wit_treess;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2388
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2389
        fun mk_coind_wit_thms ((I, dummys), (wits, wit_thms)) =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2390
          let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2391
            fun mk_goal sets y y_copy y'_copy j =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2392
              let
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2393
                fun mk_conjunct set z dummy wit =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2394
                  mk_Ball (set $ z) (Term.absfree y'_copy
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2395
                    (if dummy = NONE orelse member (op =) I (j - 1) then
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2396
                      HOLogic.mk_imp (HOLogic.mk_eq (z, wit),
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2397
                        if member (op =) I (j - 1) then HOLogic.mk_eq (y_copy, y)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2398
                        else @{term False})
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2399
                    else @{term True}));
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2400
              in
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2401
                HOLogic.mk_Trueprop
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2402
                  (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct sets Jzs dummys wits))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2403
              end;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2404
            val goals = @{map 5} mk_goal Jsetss_by_range ys ys_copy ys'_copy ls;
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2405
          in
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2406
            map2 (fn goal => fn induct =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2407
              Goal.prove_sorry lthy [] [] goal
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2408
                (fn {context = ctxt, prems = _} => mk_coind_wit_tac ctxt induct dtor_unfold_thms
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2409
                  (flat set_mapss) wit_thms)
56272
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2410
              |> Thm.close_derivation
159c07ceb18c prove theorems with fixed variables (export afterwards)
traytel
parents: 56179
diff changeset
  2411
              |> singleton (Proof_Context.export names_lthy lthy))
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2412
            goals dtor_Jset_induct_thms
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2413
            |> map split_conj_thm
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2414
            |> transpose
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2415
            |> map (map_filter (try (fn thm => thm RS bspec RS mp)))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2416
            |> curry op ~~ (map_index Library.I (map (close_wit I) wits))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2417
            |> filter (fn (_, thms) => length thms = m)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2418
          end;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2419
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2420
        val coind_wit_thms = maps mk_coind_wit_thms coind_witss;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2421
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2422
        val (wit_thmss, all_witss) =
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2423
          fold (fn ((i, wit), thms) => fn witss =>
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2424
            nth_map i (fn (thms', wits) => (thms @ thms', wit :: wits)) witss)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2425
          coind_wit_thms (map (pair []) ctor_witss)
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2426
          |> map (apsnd (map snd o minimize_wits))
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2427
          |> split_list;
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2428
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2429
        val timer = time (timer "witnesses");
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2430
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2431
        val map_id0_tacs =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2432
          map2 (fn thm => fn thm' => fn ctxt =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2433
            mk_map_id0_tac ctxt Jmap_thms thm thm')
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2434
          dtor_unfold_unique_thms unfold_dtor_thms;
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2435
        val map_comp0_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS sym) 1) Jmap_comp0_thms;
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  2436
        val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) map_cong0_thms;
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  2437
        val set_map0_tacss =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2438
          map (map (fn col => fn ctxt =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2439
            unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac ctxt col))
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2440
          (transpose col_natural_thmss);
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2441
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2442
        val Jbd_card_orders = map (fn def => fold_thms lthy [def] sbd_card_order) Jbd_defs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2443
        val Jbd_Cinfinites = map (fn def => fold_thms lthy [def] sbd_Cinfinite) Jbd_defs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2444
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2445
        val bd_co_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) Jbd_card_orders;
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2446
        val bd_cinf_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS conjunct1) 1) Jbd_Cinfinites;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2447
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2448
        val set_bd_tacss =
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  2449
          map2 (fn Cinf => map (fn col => fn ctxt =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2450
            unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac ctxt Cinf col))
56113
e3b8f8319d73 simplified internal codatatype construction
traytel
parents: 56017
diff changeset
  2451
          Jbd_Cinfinites (transpose col_bd_thmss);
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2452
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2453
        val le_rel_OO_tacs = map (fn i => fn ctxt =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2454
          rtac ctxt (le_Jrel_OO_thm RS mk_conjunctN n i) 1) ks;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2455
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59936
diff changeset
  2456
        val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Jrel_unabs_defs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2457
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2458
        val tacss = @{map 9} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2459
          bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2460
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2461
        fun wit_tac thms ctxt =
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
  2462
          mk_wit_tac ctxt n dtor_ctor_thms (flat dtor_Jset_thmss) (maps wit_thms_of_bnf bnfs) thms;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2463
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2464
        val (Jbnfs, lthy) =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58585
diff changeset
  2465
          @{fold_map 6} (fn tacs => fn map_b => fn rel_b => fn set_bs => fn wit_thms =>
56348
blanchet
parents: 56346
diff changeset
  2466
              fn consts =>
56179
6b5c46582260 tuned internal construction
traytel
parents: 56114
diff changeset
  2467
            bnf_def Hardly_Inline (user_policy Note_Some) false I tacs (wit_tac wit_thms)
58177
166131276380 introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents: 57991
diff changeset
  2468
              (SOME deads) map_b rel_b set_bs consts)
58179
2de7b0313de3 tuned size function generation
blanchet
parents: 58177
diff changeset
  2469
          tacss map_bs rel_bs set_bss wit_thmss
56766
ba2fa4e99729 more reliable 'name_of_bnf'
blanchet
parents: 56516
diff changeset
  2470
          ((((((replicate n Binding.empty ~~ Ts) ~~ Jmaps) ~~ Jsetss_by_bnf) ~~ Jbds) ~~
ba2fa4e99729 more reliable 'name_of_bnf'
blanchet
parents: 56516
diff changeset
  2471
            all_witss) ~~ map SOME Jrels)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2472
          lthy;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2473
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2474
        val timer = time (timer "registered new codatatypes as BNFs");
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2475
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2476
        val ls' = if m = 1 then [0] else ls;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2477
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2478
        val Jbnf_common_notes =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2479
          map2 (fn i => fn thm => (mk_dtor_set_inductN i, [thm])) ls' dtor_Jset_induct_thms
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2480
          |> map (fn (thmN, thms) =>
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2481
            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2482
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2483
        val Jbnf_notes =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2484
          [(dtor_mapN, map single dtor_Jmap_thms),
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2485
          (dtor_map_uniqueN, map single dtor_Jmap_unique_thms),
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2486
          (dtor_relN, map single dtor_Jrel_thms),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2487
          (dtor_set_inclN, dtor_Jset_incl_thmss),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2488
          (dtor_set_set_inclN, map flat dtor_set_Jset_incl_thmsss)] @
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2489
          map2 (fn i => fn thms => (mk_dtor_setN i, map single thms)) ls' dtor_Jset_thmss
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2490
          |> maps (fn (thmN, thmss) =>
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2491
            map2 (fn b => fn thms =>
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2492
              ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2493
            bs thmss)
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2494
      in
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2495
        (timer, Jbnfs, (Jmap_thms, dtor_Jmap_thms), dtor_Jmap_unique_thms, dtor_Jset_thmss',
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57631
diff changeset
  2496
          dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, dtor_Jset_induct_thms,
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57631
diff changeset
  2497
          lthy)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2498
      end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2499
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2500
    val dtor_unfold_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2501
      dtor_unfold_unique_thm dtor_Jmap_o_thms (map (mk_pointfree lthy) dtor_unfold_thms)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2502
      sym_map_comps map_cong0s;
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2503
    val dtor_corec_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP true m
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2504
      dtor_corec_unique_thm dtor_Jmap_o_thms (map (mk_pointfree lthy) dtor_corec_thms)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2505
      sym_map_comps map_cong0s;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2506
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2507
    val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
58443
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2508
    val Jrels = if m = 0 then map HOLogic.eq_const Ts
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2509
      else map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54793
diff changeset
  2510
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2511
    val dtor_unfold_transfer_thms =
58443
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2512
      mk_co_iter_transfer_thms Greatest_FP rels activephis activephis Jrels Jphis
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2513
        (mk_unfolds passiveAs activeAs) (mk_unfolds passiveBs activeBs)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2514
        (fn {context = ctxt, prems = _} => mk_unfold_transfer_tac ctxt m Jrel_coinduct_thm
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2515
          (map map_transfer_of_bnf bnfs) dtor_unfold_thms)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2516
        lthy;
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52659
diff changeset
  2517
58443
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2518
    val corec_allAs = passiveAs @ map2 (curry mk_sumT) Ts activeAs;
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2519
    val corec_allBs = passiveBs @ map2 (curry mk_sumT) Ts' activeBs;
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2520
    val corec_rels = map2 (fn Ds => mk_rel_of_bnf Ds corec_allAs corec_allBs) Dss bnfs;
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2521
    val corec_activephis =
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2522
      map2 (fn Jrel => mk_rel_sum (Term.list_comb (Jrel, Jphis))) Jrels activephis;
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2523
    val dtor_corec_transfer_thms =
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2524
      mk_co_iter_transfer_thms Greatest_FP corec_rels corec_activephis activephis Jrels Jphis
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2525
        (mk_corecs Ts passiveAs activeAs) (mk_corecs Ts' passiveBs activeBs)
58445
86b5b02ef33a generate 'dtor_corec_transfer' for codatatypes
desharna
parents: 58443
diff changeset
  2526
        (fn {context = ctxt, prems = _} => mk_dtor_corec_transfer_tac ctxt n m corec_defs
86b5b02ef33a generate 'dtor_corec_transfer' for codatatypes
desharna
parents: 58443
diff changeset
  2527
           dtor_unfold_transfer_thms (map map_transfer_of_bnf bnfs) dtor_Jrel_thms)
58443
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2528
        lthy;
a23780c22245 goal generation for xtor_co_rec_transfer
traytel
parents: 58256
diff changeset
  2529
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2530
    val timer = time (timer "relator coinduction");
51925
e3b7917186f1 relator coinduction for codatatypes
traytel
parents: 51917
diff changeset
  2531
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2532
    val common_notes =
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2533
      [(dtor_coinductN, [dtor_coinduct_thm]),
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2534
      (dtor_rel_coinductN, [Jrel_coinduct_thm])]
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2535
      |> map (fn (thmN, thms) =>
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2536
        ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49105
diff changeset
  2537
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2538
    val notes =
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2539
      [(ctor_dtorN, ctor_dtor_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2540
      (ctor_exhaustN, ctor_exhaust_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2541
      (ctor_injectN, ctor_inject_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2542
      (dtor_corecN, dtor_corec_thms),
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2543
      (dtor_corec_o_mapN, dtor_corec_o_Jmap_thms),
58445
86b5b02ef33a generate 'dtor_corec_transfer' for codatatypes
desharna
parents: 58443
diff changeset
  2544
      (dtor_corec_transferN, dtor_corec_transfer_thms),
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2545
      (dtor_corec_uniqueN, dtor_corec_unique_thms),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2546
      (dtor_ctorN, dtor_ctor_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2547
      (dtor_exhaustN, dtor_exhaust_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2548
      (dtor_injectN, dtor_inject_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2549
      (dtor_unfoldN, dtor_unfold_thms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2550
      (dtor_unfold_o_mapN, dtor_unfold_o_Jmap_thms),
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2551
      (dtor_unfold_transferN, dtor_unfold_transfer_thms),
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2552
      (dtor_unfold_uniqueN, dtor_unfold_unique_thms)]
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2553
      |> map (apsnd (map single))
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2554
      |> maps (fn (thmN, thmss) =>
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2555
        map2 (fn b => fn thms =>
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2556
          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2557
        bs thmss);
53568
f9456284048f conceal low-level noted facts (+ FIXME to get rid of the notes altogether eventually)
traytel
parents: 53567
diff changeset
  2558
57991
f50b3726266f don't note low-level (co)datatype theorems, unless 'bnf_note_all' is set
blanchet
parents: 57932
diff changeset
  2559
    val lthy' = lthy |> note_all ? snd o Local_Theory.notes (common_notes @ notes @ Jbnf_notes);
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2560
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2561
    val fp_res =
59819
dbec7f33093d store low-level (un)fold constants
blanchet
parents: 59621
diff changeset
  2562
      {Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors, xtor_un_folds = unfolds,
dbec7f33093d store low-level (un)fold constants
blanchet
parents: 59621
diff changeset
  2563
       xtor_co_recs = corecs, xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms,
dbec7f33093d store low-level (un)fold constants
blanchet
parents: 59621
diff changeset
  2564
       ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms,
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2565
       dtor_injects = dtor_inject_thms, xtor_maps = dtor_Jmap_thms,
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2566
       xtor_map_uniques = dtor_Jmap_unique_thms, xtor_setss = dtor_Jset_thmss',
59819
dbec7f33093d store low-level (un)fold constants
blanchet
parents: 59621
diff changeset
  2567
       xtor_rels = dtor_Jrel_thms, xtor_un_fold_thms = dtor_unfold_thms,
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2568
       xtor_co_rec_thms = dtor_corec_thms, xtor_un_fold_uniques = dtor_unfold_unique_thms,
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2569
       xtor_co_rec_uniques = dtor_corec_unique_thms, xtor_un_fold_o_maps = dtor_unfold_o_Jmap_thms,
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2570
       xtor_co_rec_o_maps = dtor_corec_o_Jmap_thms,
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2571
       xtor_un_fold_transfers = dtor_unfold_transfer_thms,
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2572
       xtor_co_rec_transfers = dtor_corec_transfer_thms, xtor_rel_co_induct = Jrel_coinduct_thm,
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59819
diff changeset
  2573
       dtor_set_inducts = dtor_Jset_induct_thms};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2574
  in
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2575
    timer; (fp_res, lthy')
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2576
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2577
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2578
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59860
diff changeset
  2579
  Outer_Syntax.local_theory @{command_keyword codatatype} "define coinductive datatypes"
52207
21026c312cc3 tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents: 52090
diff changeset
  2580
    (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
  2581
58256
08c0f0d4b9f4 generalized 'datatype' LaTeX antiquotation and added 'codatatype'
blanchet
parents: 58241
diff changeset
  2582
val _ = Theory.setup (fp_antiquote_setup @{binding codatatype});
08c0f0d4b9f4 generalized 'datatype' LaTeX antiquotation and added 'codatatype'
blanchet
parents: 58241
diff changeset
  2583
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2584
end;