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