src/HOL/BNF/Tools/bnf_lfp_tactics.ML
author traytel
Thu, 25 Apr 2013 19:18:20 +0200
changeset 51782 84e7225f5ab6
parent 51766 f19a4d0ab1bf
child 51798 ad3a241def73
permissions -rw-r--r--
removed unnecessary assumptions in some theorems about cardinal exponentiation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49506
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_lfp_tactics.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
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
Tactics for the datatype construction.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
signature BNF_LFP_TACTICS =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
sig
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
  val mk_alg_min_alg_tac: int -> thm -> thm list -> thm -> thm -> thm list list -> thm list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
    thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
  val mk_alg_not_empty_tac: thm -> thm list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
  val mk_alg_select_tac: thm -> {prems: 'a, context: Proof.context} -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
  val mk_alg_set_tac: thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
  val mk_bd_card_order_tac: thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  val mk_bd_limit_tac: int -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
  val mk_card_of_min_alg_tac: thm -> thm -> thm -> thm -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
  val mk_copy_alg_tac: thm list list -> thm list -> thm -> thm -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
  val mk_copy_str_tac: thm list list -> thm -> thm list -> tactic
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
    21
  val mk_ctor_induct_tac: int -> thm list list -> thm -> thm list -> thm -> thm list -> thm list ->
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
    22
    thm list -> tactic
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
    23
  val mk_ctor_induct2_tac: ctyp option list -> cterm option list -> thm -> thm list ->
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
    {prems: 'a, context: Proof.context} -> tactic
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
    25
  val mk_ctor_set_tac: thm -> thm -> thm list -> tactic
49518
b377da40244b renamed LFP low-level rel property to have ctor not dtor in its name
blanchet
parents: 49516
diff changeset
    26
  val mk_ctor_srel_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm -> thm ->
b377da40244b renamed LFP low-level rel property to have ctor not dtor in its name
blanchet
parents: 49516
diff changeset
    27
    thm list -> thm list -> thm list list -> tactic
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
    28
  val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list -> tactic
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
    29
  val mk_ex_copy_alg_tac: int -> thm -> thm -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
  val mk_in_bd_tac: thm -> thm -> thm -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
  val mk_incl_min_alg_tac: (int -> tactic) -> thm list list list -> thm list -> thm ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
    {prems: 'a, context: Proof.context} -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
  val mk_init_ex_mor_tac: thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
    {prems: 'a, context: Proof.context} -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
  val mk_init_induct_tac: int -> thm -> thm -> thm list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
  val mk_init_unique_mor_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
    thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
  val mk_iso_alt_tac: thm list -> thm -> tactic
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
    39
  val mk_fold_unique_mor_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
  val mk_least_min_alg_tac: thm -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  val mk_lfp_map_wpull_tac: int -> (int -> tactic) -> thm list -> thm list -> thm list list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
    thm list list list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
  val mk_map_comp_tac: thm list -> thm list -> thm -> int -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
  val mk_map_id_tac: thm list -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
  val mk_map_tac: int -> int -> thm -> thm -> thm -> tactic
49543
53b3c532a082 renamed low-level "map_unique" to have "ctor" or "dtor" in the name
blanchet
parents: 49542
diff changeset
    46
  val mk_ctor_map_unique_tac: int -> thm -> thm -> thm list -> thm list -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
  val mk_mcong_tac: (int -> tactic) -> thm list list list -> thm list -> thm list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
    {prems: 'a, context: Proof.context} -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
  val mk_min_algs_card_of_tac: ctyp -> cterm -> int -> thm -> thm list -> thm list -> thm -> thm ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
    thm -> thm -> thm -> thm -> thm -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
  val mk_min_algs_least_tac: ctyp -> cterm -> thm -> thm list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
  val mk_min_algs_mono_tac: thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
  val mk_min_algs_tac: thm -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
  val mk_mor_Abs_tac: thm -> thm list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
  val mk_mor_Rep_tac: thm list -> thm -> thm list -> thm list -> thm list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
    {prems: 'a, context: Proof.context} -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
  val mk_mor_UNIV_tac: int -> thm list -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
  val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
  val mk_mor_convol_tac: 'a list -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
  val mk_mor_elim_tac: thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
  val mk_mor_incl_tac: thm -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
  val mk_mor_inv_tac: thm -> thm -> thm list list -> thm list -> thm list -> thm list -> tactic
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
    63
  val mk_mor_fold_tac: ctyp -> cterm -> thm list -> thm -> thm -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
  val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
    thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
  val mk_mor_str_tac: 'a list -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
  val mk_rec_tac: thm list -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
    68
  val mk_rec_unique_mor_tac: thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
    69
    tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
  val mk_set_bd_tac: int -> (int -> tactic) -> thm -> thm list list -> thm list -> int ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
    {prems: 'a, context: Proof.context} -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
  val mk_set_nat_tac: int -> (int -> tactic) -> thm list list -> thm list -> cterm list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
    thm list -> int -> {prems: 'a, context: Proof.context} -> tactic
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
    74
  val mk_set_map_tac: thm -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
  val mk_set_tac: thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
  val mk_wit_tac: int -> thm list -> thm list -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
  val mk_wpull_tac: thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
structure BNF_LFP_Tactics : BNF_LFP_TACTICS =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
open BNF_Tactics
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
open BNF_LFP_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
open BNF_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
    87
val fst_snd_convs = @{thms fst_conv snd_conv};
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
    88
val ord_eq_le_trans = @{thm ord_eq_le_trans};
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
    89
val subset_trans = @{thm subset_trans};
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
    90
val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
    91
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
fun mk_alg_set_tac alg_def =
49488
02eb07152998 use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
traytel
parents: 49463
diff changeset
    93
  dtac (alg_def RS iffD1) 1 THEN
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
  REPEAT_DETERM (etac conjE 1) THEN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
  EVERY' [etac bspec, rtac CollectI] 1 THEN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
  REPEAT_DETERM (etac conjI 1) THEN atac 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
fun mk_alg_not_empty_tac alg_set alg_sets wits =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
  (EVERY' [rtac notI, hyp_subst_tac, ftac alg_set] THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
  REPEAT_DETERM o FIRST'
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   101
    [rtac subset_UNIV,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
    EVERY' [rtac @{thm subset_emptyI}, eresolve_tac wits],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
    EVERY' [rtac subsetI, rtac FalseE, eresolve_tac wits],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
    EVERY' [rtac subsetI, dresolve_tac wits, hyp_subst_tac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
      FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
  etac @{thm emptyE}) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
fun mk_mor_elim_tac mor_def =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
  (dtac (subst OF [mor_def]) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
  REPEAT o etac conjE THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
  TRY o rtac @{thm image_subsetI} THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
  etac bspec THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
  atac) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
fun mk_mor_incl_tac mor_def map_id's =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
  (stac mor_def THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
  rtac conjI THEN'
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   118
  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_id's THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
  CONJ_WRAP' (fn thm =>
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   120
   (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_id's) 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   122
fun mk_mor_comp_tac mor_def set_map's map_comp_ids =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
    val fbetw_tac = EVERY' [rtac ballI, stac o_apply, etac bspec, etac bspec, atac];
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   125
    fun mor_tac (set_map', map_comp_id) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
      EVERY' [rtac ballI, stac o_apply, rtac trans,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
        rtac trans, dtac @{thm rev_bspec}, atac, etac arg_cong,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
         REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
      CONJ_WRAP' (fn thm =>
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   130
        FIRST' [rtac subset_UNIV,
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   131
          (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   132
            etac bspec, etac set_mp, atac])]) set_map' THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
      rtac (map_comp_id RS arg_cong);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
    (dtac (mor_def RS subst) THEN' dtac (mor_def RS subst) THEN' stac mor_def THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
    REPEAT o etac conjE THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
    rtac conjI THEN'
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   138
    CONJ_WRAP' (K fbetw_tac) set_map's THEN'
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   139
    CONJ_WRAP' mor_tac (set_map's ~~ map_comp_ids)) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   142
fun mk_mor_inv_tac alg_def mor_def set_map's morEs map_comp_ids map_cong0Ls =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
  let
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   144
    val fbetw_tac = EVERY' [rtac ballI, etac set_mp, etac imageI];
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   145
    fun Collect_tac set_map' =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
      CONJ_WRAP' (fn thm =>
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   147
        FIRST' [rtac subset_UNIV,
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   148
          (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   149
            etac @{thm image_mono}, atac])]) set_map';
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   150
    fun mor_tac (set_map', ((morE, map_comp_id), map_cong0L)) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
      EVERY' [rtac ballI, ftac @{thm rev_bspec}, atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
         REPEAT o eresolve_tac [CollectE, conjE], rtac sym, rtac trans, rtac sym,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   153
         etac @{thm inverE}, etac bspec, rtac CollectI, Collect_tac set_map',
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   154
         rtac trans, etac (morE RS arg_cong), rtac CollectI, Collect_tac set_map',
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   155
         rtac trans, rtac (map_comp_id RS arg_cong), rtac (map_cong0L RS arg_cong),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
         REPEAT_DETERM_N (length morEs) o
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
           (EVERY' [rtac subst, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
    (stac mor_def THEN'
49488
02eb07152998 use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
traytel
parents: 49463
diff changeset
   160
    dtac (alg_def RS iffD1) THEN'
02eb07152998 use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
traytel
parents: 49463
diff changeset
   161
    dtac (alg_def RS iffD1) THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
    REPEAT o etac conjE THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
    rtac conjI THEN'
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   164
    CONJ_WRAP' (K fbetw_tac) set_map's THEN'
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   165
    CONJ_WRAP' mor_tac (set_map's ~~ (morEs ~~ map_comp_ids ~~ map_cong0Ls))) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
fun mk_mor_str_tac ks mor_def =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
  (stac mor_def THEN' rtac conjI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
fun mk_mor_convol_tac ks mor_def =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
  (stac mor_def THEN' rtac conjI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1;
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
fun mk_mor_UNIV_tac m morEs mor_def =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
    val n = length morEs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
    fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   182
      rtac CollectI, CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto m + n),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
      rtac sym, rtac o_apply];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
    EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
    REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS subst),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
    CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
      etac (o_apply RS subst), rtac o_apply])) morEs] 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
fun mk_iso_alt_tac mor_images mor_inv =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   194
    val n = length mor_images;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   195
    fun if_wrap_tac thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
      EVERY' [rtac ssubst, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   197
        rtac @{thm inver_surj}, etac thm, etac thm, atac, etac conjI, atac]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
    val if_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
      EVERY' [etac thin_rl, etac thin_rl, REPEAT o eresolve_tac [conjE, exE],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
        rtac conjI, atac, CONJ_WRAP' if_wrap_tac mor_images];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
    val only_if_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
      EVERY' [rtac conjI, etac conjunct1, EVERY' (map (fn thm =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   203
        EVERY' [rtac exE, rtac @{thm bij_betw_ex_weakE}, etac (conjunct2 RS thm)])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
        (map (mk_conjunctN n) (1 upto n))), REPEAT o rtac exI, rtac conjI, rtac mor_inv,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
        etac conjunct1, atac, atac, REPEAT_DETERM_N n o atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
        CONJ_WRAP' (K (etac conjunct2)) mor_images];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
    (rtac iffI THEN' if_tac THEN' only_if_tac) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   211
fun mk_copy_str_tac set_map's alg_def alg_sets =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   213
    val n = length alg_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   214
    val bij_betw_inv_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   215
      EVERY' [etac thin_rl, REPEAT_DETERM_N n o EVERY' [dtac @{thm bij_betwI}, atac, atac],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
        REPEAT_DETERM_N (2 * n) o etac thin_rl, REPEAT_DETERM_N (n - 1) o etac conjI, atac];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   217
    fun set_tac thms =
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   218
      EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
          etac @{thm image_mono}, rtac equalityD1, etac @{thm bij_betw_imageE}];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
    val copy_str_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   221
      CONJ_WRAP' (fn (thms, thm) =>
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   222
        EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac set_mp,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   223
          rtac equalityD1, etac @{thm bij_betw_imageE}, rtac imageI, etac thm,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   224
          REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   225
      (set_map's ~~ alg_sets);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   226
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
    (rtac rev_mp THEN' DETERM o bij_betw_inv_tac THEN' rtac impI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
    stac alg_def THEN' copy_str_tac) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   231
fun mk_copy_alg_tac set_map's alg_sets mor_def iso_alt copy_str =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
    val n = length alg_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
    val fbetw_tac = CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
    fun set_tac thms =
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   236
      EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
        REPEAT_DETERM o etac conjE, etac @{thm image_mono},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   238
        rtac equalityD1, etac @{thm bij_betw_imageE}];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
    val mor_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
      CONJ_WRAP' (fn (thms, thm) =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
        EVERY' [rtac ballI, etac CollectE, etac @{thm inverE}, etac thm,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   242
          REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   243
      (set_map's ~~ alg_sets);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
    (rtac (iso_alt RS @{thm ssubst[of _ _ "%x. x"]}) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
    etac copy_str THEN' REPEAT_DETERM o atac THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   247
    rtac conjI THEN' stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
    CONJ_WRAP' (K atac) alg_sets) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
fun mk_ex_copy_alg_tac n copy_str copy_alg =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   252
  EVERY' [REPEAT_DETERM_N n o rtac exI, rtac conjI, etac copy_str,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
    REPEAT_DETERM_N n o atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   254
    REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
    REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}, etac copy_alg,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
    REPEAT_DETERM_N n o atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
    REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
    REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}] 1;
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
fun mk_bd_limit_tac n bd_Cinfinite =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
  EVERY' [REPEAT_DETERM o etac conjE, rtac rev_mp, rtac @{thm Cinfinite_limit_finite},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
    REPEAT_DETERM_N n o rtac @{thm finite.insertI}, rtac @{thm finite.emptyI},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
    REPEAT_DETERM_N n o etac @{thm insert_subsetI}, rtac @{thm empty_subsetI},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
    rtac bd_Cinfinite, rtac impI, etac bexE, rtac bexI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
    CONJ_WRAP' (fn i =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
      EVERY' [etac bspec, REPEAT_DETERM_N i o rtac @{thm insertI2}, rtac @{thm insertI1}])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
      (0 upto n - 1),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
    atac] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
fun mk_min_algs_tac worel in_congs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
    val minG_tac = EVERY' [rtac @{thm UN_cong}, rtac refl, dtac bspec, atac, etac arg_cong];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
    fun minH_tac thm =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   274
      EVERY' [rtac @{thm Un_cong}, minG_tac, rtac @{thm image_cong}, rtac thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   275
        REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   276
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
    (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ssubst THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
    rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
    REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN'
49488
02eb07152998 use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
traytel
parents: 49463
diff changeset
   280
    CONJ_WRAP_GEN' (EVERY' [rtac Pair_eqI, rtac conjI]) minH_tac in_congs) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
fun mk_min_algs_mono_tac min_algs = EVERY' [stac @{thm relChain_def}, rtac allI, rtac allI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
  rtac impI, rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
  rtac subsetI, rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
  rtac equalityD1, dtac @{thm notnotD}, hyp_subst_tac, rtac refl] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
  suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite Asuc_Cnotzero =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
    val induct = worel RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
      Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
    val src = 1 upto m + 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
    val dest = (m + 1) :: (1 upto m);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
    val absorbAs_tac = if m = 0 then K (all_tac)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
      else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   297
        rtac @{thm ordIso_transitive},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   298
        BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   299
          FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
            rtac @{thm Card_order_cexp}])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   301
        @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   302
        src dest,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
        rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac ctrans, rtac @{thm ordLeq_csum1},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
        FIRST' [rtac @{thm Card_order_csum}, rtac @{thm card_of_Card_order}],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
        rtac @{thm ordLeq_cexp1}, rtac suc_Cnotzero, rtac @{thm Card_order_csum}];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   307
    val minG_tac = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordLess_imp_ordLeq},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   308
      rtac @{thm ordLess_transitive}, rtac @{thm card_of_underS}, rtac suc_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   309
      atac, rtac suc_Asuc, rtac ballI, etac allE, dtac mp, etac @{thm underS_E},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   310
      dtac mp, etac @{thm underS_Field}, REPEAT o etac conjE, atac, rtac Asuc_Cinfinite]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   311
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   312
    fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac @{thm ordIso_ordLeq_trans},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   313
      rtac @{thm card_of_ordIso_subst}, etac min_alg, rtac @{thm Un_Cinfinite_bound},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   314
      minG_tac, rtac ctrans, rtac @{thm card_of_image}, rtac ctrans, rtac in_bd, rtac ctrans,
51782
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   315
      rtac @{thm cexp_mono1}, rtac @{thm csum_mono1},
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
      REPEAT_DETERM_N m o rtac @{thm csum_mono2},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   317
      CONJ_WRAP_GEN' (rtac @{thm csum_cinfinite_bound}) (K minG_tac) min_algs,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   318
      REPEAT_DETERM o FIRST'
51782
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   319
        [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   320
        rtac Asuc_Cinfinite, rtac bd_Card_order],
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   321
      rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1}, absorbAs_tac,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   322
      rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
      rtac Asuc_Cinfinite, rtac bd_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
      rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq},
51782
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   325
      resolve_tac @{thms Card_order_csum Card_order_ctwo}, rtac suc_Cinfinite,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
      rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   328
    (rtac induct THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   329
    rtac impI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
    CONJ_WRAP' mk_minH_tac (min_algs ~~ in_bds)) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   331
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   332
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   333
fun mk_min_algs_least_tac cT ct worel min_algs alg_sets =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   334
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   335
    val induct = worel RS
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   336
      Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   337
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   338
    val minG_tac = EVERY' [rtac @{thm UN_least}, etac allE, dtac mp, etac @{thm underS_E},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   339
      dtac mp, etac @{thm underS_Field}, REPEAT_DETERM o etac conjE, atac];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   340
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   341
    fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   342
      rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
      REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac alg_set,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
      REPEAT_DETERM o FIRST' [atac, etac subset_trans THEN' minG_tac]];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   345
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   346
    (rtac induct THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   347
    rtac impI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
    CONJ_WRAP' mk_minH_tac (min_algs ~~ alg_sets)) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   349
  end;
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
fun mk_alg_min_alg_tac m alg_def min_alg_defs bd_limit bd_Cinfinite
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   352
    set_bdss min_algs min_alg_monos =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   353
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   354
    val n = length min_algs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   355
    fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   356
      [rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
       etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve_tac set_bds];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   358
    fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   359
      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   360
        EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
        rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   362
        rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}),
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   363
        rtac @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac thin_rl, atac, rtac set_mp,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   364
        rtac equalityD2, rtac min_alg, atac, rtac UnI2, rtac @{thm image_eqI}, rtac refl,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   365
        rtac CollectI, REPEAT_DETERM_N m o dtac asm_rl, REPEAT_DETERM_N n o etac thin_rl,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   366
        REPEAT_DETERM o etac conjE,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   367
        CONJ_WRAP' (K (FIRST' [atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   368
          EVERY' [etac subset_trans, rtac subsetI, rtac @{thm UN_I},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   369
            etac @{thm underS_I}, atac, atac]]))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   370
          set_bds];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   371
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   372
    (rtac (alg_def RS @{thm ssubst[of _ _ "%x. x"]}) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   373
    CONJ_WRAP' mk_conjunct_tac (set_bdss ~~ (min_algs ~~ min_alg_defs))) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   374
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   375
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   376
fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   377
  EVERY' [stac min_alg_def, rtac @{thm UNION_Cinfinite_bound},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   378
    rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
    rtac @{thm ordLess_imp_ordLeq}, rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   380
    REPEAT_DETERM o etac conjE, atac, rtac Asuc_Cinfinite] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   381
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
fun mk_least_min_alg_tac min_alg_def least =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
  EVERY' [stac min_alg_def, rtac @{thm UN_least}, dtac least, dtac mp, atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   384
    REPEAT_DETERM o etac conjE, atac] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   385
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   386
fun mk_alg_select_tac Abs_inverse {context = ctxt, prems = _} =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   387
  EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac] 1 THEN
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   388
  unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   389
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   390
fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   391
    alg_sets set_map's str_init_defs =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   392
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   393
    val n = length alg_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   394
    val fbetw_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   395
      CONJ_WRAP' (K (EVERY' [rtac ballI, etac @{thm rev_bspec}, etac CollectE, atac])) alg_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   396
    val mor_tac =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   397
      CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   398
    fun alg_epi_tac ((alg_set, str_init_def), set_map') =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   399
      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   400
        rtac ballI, ftac (alg_select RS bspec), stac str_init_def, etac alg_set,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   401
        REPEAT_DETERM o FIRST' [rtac subset_UNIV,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   402
          EVERY' [rtac ord_eq_le_trans, resolve_tac set_map', rtac subset_trans,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   403
            etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, atac]]];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   404
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   405
    (rtac mor_cong THEN' REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm o_id}) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   406
    rtac (Thm.permute_prems 0 1 mor_comp) THEN' etac (Thm.permute_prems 0 1 mor_comp) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   407
    stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN' rtac mor_incl_min_alg THEN'
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   408
    stac alg_def THEN' CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_map's)) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   411
fun mk_init_ex_mor_tac Abs_inverse copy_alg_ex alg_min_alg card_of_min_algs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   412
    mor_comp mor_select mor_incl_min_alg {context = ctxt, prems = _} =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   413
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   414
    val n = length card_of_min_algs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   415
    val card_of_ordIso_tac = EVERY' [rtac ssubst, rtac @{thm card_of_ordIso},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   416
      rtac @{thm ordIso_symmetric}, rtac conjunct1, rtac conjunct2, atac];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   417
    fun internalize_tac card_of = EVERY' [rtac subst, rtac @{thm internalize_card_of_ordLeq2},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
      rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac subst,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   419
      rtac @{thm Card_order_iff_ordIso_card_of}, rtac @{thm Card_order_cexp}];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   420
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   421
    (rtac rev_mp THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
    REPEAT_DETERM_N (2 * n) o (rtac mp THEN' rtac @{thm ex_mono} THEN' rtac impI) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   423
    REPEAT_DETERM_N (n + 1) o etac thin_rl THEN' rtac (alg_min_alg RS copy_alg_ex) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   424
    REPEAT_DETERM_N n o atac THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   425
    REPEAT_DETERM_N n o card_of_ordIso_tac THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   426
    EVERY' (map internalize_tac card_of_min_algs) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   427
    rtac impI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   428
    REPEAT_DETERM o eresolve_tac [exE, conjE] THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   429
    REPEAT_DETERM o rtac exI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   430
    rtac mor_select THEN' atac THEN' rtac CollectI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   431
    REPEAT_DETERM o rtac exI THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
    rtac conjI THEN' rtac refl THEN' atac THEN'
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   433
    K (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)) THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   434
    etac mor_comp THEN' etac mor_incl_min_alg) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   435
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   436
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   437
fun mk_init_unique_mor_tac m
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   438
    alg_def alg_min_alg least_min_algs in_monos alg_sets morEs map_cong0s =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   439
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   440
    val n = length least_min_algs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   441
    val ks = (1 upto n);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   442
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   443
    fun mor_tac morE in_mono = EVERY' [etac morE, rtac set_mp, rtac in_mono,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   444
      REPEAT_DETERM_N n o rtac @{thm Collect_restrict}, rtac CollectI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   445
      REPEAT_DETERM_N (m + n) o (TRY o rtac conjI THEN' atac)];
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   446
    fun cong_tac map_cong0 = EVERY' [rtac (map_cong0 RS arg_cong),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   447
      REPEAT_DETERM_N m o rtac refl,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   448
      REPEAT_DETERM_N n o (etac @{thm prop_restrict} THEN' atac)];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   449
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   450
    fun mk_alg_tac (alg_set, (in_mono, (morE, map_cong0))) = EVERY' [rtac ballI, rtac CollectI,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   451
      REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   452
      REPEAT_DETERM_N m o rtac subset_UNIV,
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   453
      REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   454
      rtac trans, mor_tac morE in_mono,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   455
      rtac trans, cong_tac map_cong0,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   456
      rtac sym, mor_tac morE in_mono];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   457
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   458
    fun mk_unique_tac (k, least_min_alg) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   459
      select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   460
      stac alg_def THEN'
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   461
      CONJ_WRAP' mk_alg_tac (alg_sets ~~ (in_monos ~~ (morEs ~~ map_cong0s)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   462
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   463
    CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   464
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   465
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   466
fun mk_init_induct_tac m alg_def alg_min_alg least_min_algs alg_sets =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   467
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   468
    val n = length least_min_algs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   469
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   470
    fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   471
      REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   472
      REPEAT_DETERM_N m o rtac subset_UNIV,
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   473
      REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   474
      rtac mp, etac bspec, rtac CollectI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   475
      REPEAT_DETERM_N m o (rtac conjI THEN' atac),
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   476
      CONJ_WRAP' (K (etac subset_trans THEN' rtac @{thm Collect_restrict})) alg_sets,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   477
      CONJ_WRAP' (K (rtac ballI THEN' etac @{thm prop_restrict} THEN' atac)) alg_sets];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   479
    fun mk_induct_tac least_min_alg =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   480
      rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   481
      stac alg_def THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   482
      CONJ_WRAP' mk_alg_tac alg_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   483
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   484
    CONJ_WRAP' mk_induct_tac least_min_algs 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   485
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   486
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   487
fun mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps {context = ctxt, prems = _} =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   488
  (K (unfold_thms_tac ctxt ctor_defs) THEN' rtac conjunct1 THEN' rtac copy THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   489
  EVERY' (map (fn bij => EVERY' [rtac bij, atac, etac bexI, rtac UNIV_I]) bijs) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   490
  EVERY' (map rtac inver_Abss) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   491
  EVERY' (map rtac inver_Reps)) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   492
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   493
fun mk_mor_Abs_tac inv inver_Abss inver_Reps =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   494
  (rtac inv THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   495
  EVERY' (map2 (fn inver_Abs => fn inver_Rep =>
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   496
    EVERY' [rtac conjI, rtac subset_UNIV, rtac conjI, rtac inver_Rep, rtac inver_Abs])
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   497
    inver_Abss inver_Reps)) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   498
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   499
fun mk_mor_fold_tac cT ct fold_defs ex_mor mor =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   500
  (EVERY' (map stac fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   501
  REPEAT_DETERM_N (length fold_defs) o etac exE THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   502
  rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   503
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   504
fun mk_fold_unique_mor_tac type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold =
48975
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
    fun mk_unique type_def =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   507
      EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}),
49227
2652319c394e open typedef for datatypes
traytel
parents: 48975
diff changeset
   508
        rtac ballI, resolve_tac init_unique_mors,
2652319c394e open typedef for datatypes
traytel
parents: 48975
diff changeset
   509
        EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   510
        rtac mor_comp, rtac mor_Abs, atac,
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   511
        rtac mor_comp, rtac mor_Abs, rtac mor_fold];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   512
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   513
    CONJ_WRAP' mk_unique type_defs 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   514
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   515
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   516
fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_folds =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   517
  EVERY' [stac dtor_def, rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   518
    rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   519
    EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply]))
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   520
      ctor_o_folds),
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   521
    rtac sym, rtac id_apply] 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   522
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   523
fun mk_rec_tac rec_defs foldx fst_recs {context = ctxt, prems = _}=
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   524
  unfold_thms_tac ctxt
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   525
    (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd}) fst_recs) THEN
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   526
  EVERY' [rtac trans, rtac o_apply, rtac trans, rtac (foldx RS @{thm arg_cong[of _ _ snd]}),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   527
    rtac @{thm snd_convol'}] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   528
51739
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
   529
fun mk_rec_unique_mor_tac rec_defs fst_recs fold_unique_mor {context = ctxt, prems = _} =
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
   530
  unfold_thms_tac ctxt
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
   531
    (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd'}) fst_recs) THEN
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
   532
  etac fold_unique_mor 1;
3514b90d0a8b (co)rec is (just as the (un)fold) the unique morphism;
traytel
parents: 49585
diff changeset
   533
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   534
fun mk_ctor_induct_tac m set_map'ss init_induct morEs mor_Abs Rep_invs Abs_invs Reps =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   535
  let
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   536
    val n = length set_map'ss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   537
    val ks = 1 upto n;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   538
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   539
    fun mk_IH_tac Rep_inv Abs_inv set_map' =
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   540
      DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   541
        dtac set_rev_mp, rtac equalityD1, rtac set_map', etac imageE,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   542
        hyp_subst_tac, rtac (Abs_inv RS ssubst), etac set_mp, atac, atac];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   543
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   544
    fun mk_closed_tac (k, (morE, set_map's)) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   545
      EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   546
        rtac (mor_Abs RS morE RS arg_cong RS ssubst), atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   547
        REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   548
        EVERY' (map3 mk_IH_tac Rep_invs Abs_invs (drop m set_map's)), atac];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   549
49227
2652319c394e open typedef for datatypes
traytel
parents: 48975
diff changeset
   550
    fun mk_induct_tac (Rep, Rep_inv) =
2652319c394e open typedef for datatypes
traytel
parents: 48975
diff changeset
   551
      EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RSN (2, bspec))];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   552
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   553
    (rtac mp THEN' rtac impI THEN'
49227
2652319c394e open typedef for datatypes
traytel
parents: 48975
diff changeset
   554
    DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   555
    rtac init_induct THEN'
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   556
    DETERM o CONJ_WRAP' mk_closed_tac (ks ~~ (morEs ~~ set_map'ss))) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   557
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   558
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   559
fun mk_ctor_induct2_tac cTs cts ctor_induct weak_ctor_inducts {context = ctxt, prems = _} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   560
  let
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   561
    val n = length weak_ctor_inducts;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   562
    val ks = 1 upto n;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   563
    fun mk_inner_induct_tac induct i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   564
      EVERY' [rtac allI, fo_rtac induct ctxt,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   565
        select_prem_tac n (dtac @{thm meta_spec2}) i,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   566
        REPEAT_DETERM_N n o
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   567
          EVERY' [dtac meta_mp THEN_ALL_NEW Goal.norm_hhf_tac,
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   568
            REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS meta_mp), atac],
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   569
        atac];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   570
  in
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   571
    EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct),
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   572
      EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   573
      REPEAT_DETERM o eresolve_tac [conjE, allE],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   574
      CONJ_WRAP' (K atac) ks] 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   575
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   576
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   577
fun mk_map_tac m n foldx map_comp_id map_cong0 =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   578
  EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx, rtac trans, rtac o_apply,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   579
    rtac trans, rtac (map_comp_id RS arg_cong), rtac trans, rtac (map_cong0 RS arg_cong),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   580
    REPEAT_DETERM_N m o rtac refl,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   581
    REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply])),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   582
    rtac sym, rtac o_apply] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   583
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   584
fun mk_ctor_map_unique_tac m mor_def fold_unique_mor map_comp_ids map_cong0s =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   585
  let
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   586
    val n = length map_cong0s;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   587
    fun mk_mor (comp_id, cong) = EVERY' [rtac ballI, rtac trans, etac @{thm pointfreeE},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   588
      rtac sym, rtac trans, rtac o_apply, rtac trans, rtac (comp_id RS arg_cong),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   589
      rtac (cong RS arg_cong),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   590
      REPEAT_DETERM_N m o rtac refl,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   591
      REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply]))];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   592
  in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   593
    EVERY' [rtac fold_unique_mor, rtac ssubst, rtac mor_def, rtac conjI,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   594
      CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) map_cong0s,
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   595
      CONJ_WRAP' mk_mor (map_comp_ids ~~ map_cong0s)] 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   596
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   597
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   598
fun mk_set_tac foldx = EVERY' [rtac ext, rtac trans, rtac o_apply,
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49502
diff changeset
   599
  rtac trans, rtac foldx, rtac sym, rtac o_apply] 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   600
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   601
fun mk_ctor_set_tac set set_map' set_map's =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   602
  let
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   603
    val n = length set_map's;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   604
    fun mk_UN thm = rtac (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   605
      rtac @{thm Union_image_eq};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   606
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   607
    EVERY' [rtac (set RS @{thm pointfreeE} RS trans), rtac @{thm Un_cong},
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   608
      rtac (trans OF [set_map', trans_fun_cong_image_id_id_apply]),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   609
      REPEAT_DETERM_N (n - 1) o rtac @{thm Un_cong},
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   610
      EVERY' (map mk_UN set_map's)] 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   611
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   612
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   613
fun mk_set_nat_tac m induct_tac set_map'ss
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   614
    ctor_maps csets ctor_sets i {context = ctxt, prems = _} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   615
  let
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   616
    val n = length ctor_maps;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   617
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   618
    fun useIH set_nat = EVERY' [rtac trans, rtac @{thm image_UN}, rtac trans, rtac @{thm UN_cong},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   619
      rtac refl, Goal.assume_rule_tac ctxt, rtac sym, rtac trans, rtac @{thm UN_cong},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   620
      rtac set_nat, rtac refl, rtac @{thm UN_simps(10)}];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   621
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   622
    fun mk_set_nat cset ctor_map ctor_set set_nats =
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   623
      EVERY' [rtac trans, rtac @{thm image_cong}, rtac ctor_set, rtac refl,
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   624
        rtac sym, rtac (trans OF [ctor_map RS HOL_arg_cong cset, ctor_set RS trans]),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   625
        rtac sym, EVERY' (map rtac (trans :: @{thms image_Un Un_cong})),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   626
        rtac sym, rtac (nth set_nats (i - 1)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   627
        REPEAT_DETERM_N (n - 1) o EVERY' (map rtac (trans :: @{thms image_Un Un_cong})),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   628
        EVERY' (map useIH (drop m set_nats))];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   629
  in
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   630
    (induct_tac THEN' EVERY' (map4 mk_set_nat csets ctor_maps ctor_sets set_map'ss)) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   631
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   632
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   633
fun mk_set_bd_tac m induct_tac bd_Cinfinite set_bdss ctor_sets i {context = ctxt, prems = _} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   634
  let
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   635
    val n = length ctor_sets;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   636
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   637
    fun useIH set_bd = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac set_bd, rtac ballI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   638
      Goal.assume_rule_tac ctxt, rtac bd_Cinfinite];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   639
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   640
    fun mk_set_nat ctor_set set_bds =
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   641
      EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_ordIso_subst}, rtac ctor_set,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   642
        rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_bds (i - 1)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   643
        REPEAT_DETERM_N (n - 1) o rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   644
        EVERY' (map useIH (drop m set_bds))];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   645
  in
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   646
    (induct_tac THEN' EVERY' (map2 mk_set_nat ctor_sets set_bdss)) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   647
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   648
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   649
fun mk_mcong_tac induct_tac set_setsss map_cong0s ctor_maps {context = ctxt, prems = _} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   650
  let
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   651
    fun use_asm thm = EVERY' [etac bspec, etac set_rev_mp, rtac thm];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   652
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   653
    fun useIH set_sets = EVERY' [rtac mp, Goal.assume_rule_tac ctxt,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   654
      CONJ_WRAP' (fn thm =>
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   655
        EVERY' [rtac ballI, etac bspec, etac set_rev_mp, etac thm]) set_sets];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   656
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   657
    fun mk_map_cong0 ctor_map map_cong0 set_setss =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   658
      EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   659
        rtac trans, rtac ctor_map, rtac trans, rtac (map_cong0 RS arg_cong),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   660
        EVERY' (map use_asm (map hd set_setss)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   661
        EVERY' (map useIH (transpose (map tl set_setss))),
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   662
        rtac sym, rtac ctor_map];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   663
  in
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   664
    (induct_tac THEN' EVERY' (map3 mk_map_cong0 ctor_maps map_cong0s set_setsss)) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   665
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   666
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   667
fun mk_incl_min_alg_tac induct_tac set_setsss alg_sets alg_min_alg {context = ctxt, prems = _} =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   668
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   669
    fun use_asm thm = etac (thm RS subset_trans);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   670
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   671
    fun useIH set_sets = EVERY' [rtac subsetI, rtac mp, Goal.assume_rule_tac ctxt,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   672
      rtac CollectI, CONJ_WRAP' (fn thm => EVERY' [etac (thm RS subset_trans), atac]) set_sets];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   673
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   674
    fun mk_incl alg_set set_setss =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   675
      EVERY' [rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   676
        rtac (alg_min_alg RS alg_set),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   677
        EVERY' (map use_asm (map hd set_setss)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   678
        EVERY' (map useIH (transpose (map tl set_setss)))];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   679
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   680
    (induct_tac THEN' EVERY' (map2 mk_incl alg_sets set_setsss)) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   681
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   682
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   683
fun mk_lfp_map_wpull_tac m induct_tac wpulls ctor_maps ctor_setss set_setsss ctor_injects =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   684
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   685
    val n = length wpulls;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   686
    val ks = 1 upto n;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   687
    val ls = 1 upto m;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   688
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   689
    fun use_pass_asm thm = rtac conjI THEN' etac (thm RS subset_trans);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   690
    fun use_act_asm thm = etac (thm RS subset_trans) THEN' atac;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   691
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   692
    fun useIH set_sets i = EVERY' [rtac ssubst, rtac @{thm wpull_def},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   693
       REPEAT_DETERM_N m o etac thin_rl, select_prem_tac n (dtac asm_rl) i,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   694
       rtac allI, rtac allI, rtac impI, REPEAT_DETERM o etac conjE,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   695
       REPEAT_DETERM o dtac @{thm meta_spec},
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   696
       dtac meta_mp, atac,
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   697
       dtac meta_mp, atac, etac mp,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   698
       rtac conjI, rtac CollectI, CONJ_WRAP' use_act_asm set_sets,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   699
       rtac conjI, rtac CollectI, CONJ_WRAP' use_act_asm set_sets,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   700
       atac];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   701
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   702
    fun mk_subset thm = EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm Un_least}, atac,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   703
      REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   704
      REPEAT_DETERM_N n o
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   705
        EVERY' [rtac @{thm UN_least}, rtac CollectE, etac set_rev_mp, atac,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   706
          REPEAT_DETERM o etac conjE, atac]];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   707
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   708
    fun mk_wpull wpull ctor_map ctor_sets set_setss ctor_inject =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   709
      EVERY' [rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   710
        rtac rev_mp, rtac wpull,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   711
        EVERY' (map (fn i => REPEAT_DETERM_N (i - 1) o etac thin_rl THEN' atac) ls),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   712
        EVERY' (map2 useIH (transpose (map tl set_setss)) ks),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   713
        rtac impI, REPEAT_DETERM_N (m + n) o etac thin_rl,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   714
        dtac @{thm subst[OF wpull_def, of "%x. x"]}, etac allE, etac allE, etac impE,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   715
        rtac conjI, rtac CollectI, EVERY' (map (use_pass_asm o hd) set_setss),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   716
          CONJ_WRAP' (K (rtac subset_refl)) ks,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   717
        rtac conjI, rtac CollectI, EVERY' (map (use_pass_asm o hd) set_setss),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   718
          CONJ_WRAP' (K (rtac subset_refl)) ks,
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   719
        rtac subst, rtac ctor_inject, rtac trans, rtac sym, rtac ctor_map,
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   720
        rtac trans, atac, rtac ctor_map, REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE],
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   721
        hyp_subst_tac, rtac bexI, rtac conjI, rtac ctor_map, rtac ctor_map, rtac CollectI,
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   722
        CONJ_WRAP' mk_subset ctor_sets];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   723
  in
49542
b39354db8629 renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
blanchet
parents: 49541
diff changeset
   724
    (induct_tac THEN' EVERY' (map5 mk_wpull wpulls ctor_maps ctor_setss set_setsss ctor_injects)) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   725
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   726
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   727
(* BNF tactics *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   728
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   729
fun mk_map_id_tac map_ids unique =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   730
  (rtac sym THEN' rtac unique THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   731
  EVERY' (map (fn thm =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   732
    EVERY' [rtac trans, rtac @{thm id_o}, rtac trans, rtac sym, rtac @{thm o_id},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   733
      rtac (thm RS sym RS arg_cong)]) map_ids)) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   734
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   735
fun mk_map_comp_tac map_comps ctor_maps unique iplus1 =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   736
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   737
    val i = iplus1 - 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   738
    val unique' = Thm.permute_prems 0 i unique;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   739
    val map_comps' = drop i map_comps @ take i map_comps;
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   740
    val ctor_maps' = drop i ctor_maps @ take i ctor_maps;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   741
    fun mk_comp comp simp =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   742
      EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac o_apply,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   743
        rtac trans, rtac (simp RS arg_cong), rtac trans, rtac simp,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   744
        rtac trans, rtac (comp RS arg_cong), rtac sym, rtac o_apply];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   745
  in
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   746
    (rtac sym THEN' rtac unique' THEN' EVERY' (map2 mk_comp map_comps' ctor_maps')) 1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   747
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   748
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   749
fun mk_set_map_tac set_nat =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   750
  EVERY' (map rtac [ext, trans, o_apply, sym, trans, o_apply, set_nat]) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   751
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   752
fun mk_in_bd_tac sum_Card_order sucbd_Cnotzero incl card_of_min_alg =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   753
  EVERY' [rtac ctrans, rtac @{thm card_of_mono1}, rtac subsetI, etac rev_mp,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   754
    rtac incl, rtac ctrans, rtac card_of_min_alg, rtac @{thm cexp_mono2_Cnotzero},
51782
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   755
    rtac @{thm cardSuc_ordLeq_cpow}, rtac sum_Card_order, rtac @{thm Card_order_csum},
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   756
    rtac sucbd_Cnotzero] 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   757
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   758
fun mk_bd_card_order_tac bd_card_orders =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   759
  (rtac @{thm card_order_cpow} THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   760
    CONJ_WRAP_GEN' (rtac @{thm card_order_csum}) rtac bd_card_orders) 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   761
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   762
fun mk_wpull_tac wpull =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   763
  EVERY' [rtac ssubst, rtac @{thm wpull_def}, rtac allI, rtac allI,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   764
    rtac wpull, REPEAT_DETERM o atac] 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   765
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   766
fun mk_wit_tac n ctor_set wit =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   767
  REPEAT_DETERM (atac 1 ORELSE
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   768
    EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   769
    REPEAT_DETERM o
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   770
      (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   771
        (eresolve_tac wit ORELSE'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   772
        (dresolve_tac wit THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   773
          (etac FalseE ORELSE'
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   774
          EVERY' [hyp_subst_tac, dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   775
            REPEAT_DETERM_N n o etac UnE]))))] 1);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   776
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   777
fun mk_ctor_srel_tac in_Isrels i in_srel map_comp map_cong0 ctor_map ctor_sets ctor_inject
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   778
  ctor_dtor set_maps ctor_set_incls ctor_set_set_inclss =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   779
  let
49544
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
   780
    val m = length ctor_set_incls;
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
   781
    val n = length ctor_set_set_inclss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   782
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   783
    val (passive_set_maps, active_set_maps) = chop m set_maps;
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   784
    val in_Isrel = nth in_Isrels (i - 1);
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   785
    val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans;
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   786
    val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   787
    val if_tac =
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   788
      EVERY' [dtac (in_Isrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   789
        rtac (in_srel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   790
        EVERY' (map2 (fn set_map => fn ctor_set_incl =>
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   791
          EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map,
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   792
            rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
49544
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
   793
            rtac (ctor_set_incl RS subset_trans), etac le_arg_cong_ctor_dtor])
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   794
        passive_set_maps ctor_set_incls),
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   795
        CONJ_WRAP' (fn (in_Isrel, (set_map, ctor_set_set_incls)) =>
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   796
          EVERY' [rtac ord_eq_le_trans, rtac set_map, rtac @{thm image_subsetI},
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   797
            rtac (in_Isrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   798
            CONJ_WRAP' (fn thm =>
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   799
              EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor]))
49544
24094fa47e0d renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents: 49543
diff changeset
   800
            ctor_set_set_incls,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   801
            rtac conjI, rtac refl, rtac refl])
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   802
        (in_Isrels ~~ (active_set_maps ~~ ctor_set_set_inclss)),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   803
        CONJ_WRAP' (fn conv =>
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   804
          EVERY' [rtac trans, rtac map_comp, rtac trans, rtac map_cong0,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   805
          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   806
          REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   807
          rtac (ctor_inject RS iffD1), rtac trans, rtac sym, rtac ctor_map,
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49490
diff changeset
   808
          etac eq_arg_cong_ctor_dtor])
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   809
        fst_snd_convs];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   810
    val only_if_tac =
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   811
      EVERY' [dtac (in_srel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   812
        rtac (in_Isrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   813
        CONJ_WRAP' (fn (ctor_set, passive_set_map) =>
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49544
diff changeset
   814
          EVERY' [rtac ord_eq_le_trans, rtac ctor_set, rtac @{thm Un_least},
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   815
            rtac ord_eq_le_trans, rtac @{thm box_equals[OF _ refl]},
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   816
            rtac passive_set_map, rtac trans_fun_cong_image_id_id_apply, atac,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   817
            CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   818
              (fn (active_set_map, in_Isrel) => EVERY' [rtac ord_eq_le_trans,
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   819
                rtac @{thm UN_cong[OF _ refl]}, rtac active_set_map, rtac @{thm UN_least},
49306
c13fff97a8df tuning annotations
blanchet
parents: 49227
diff changeset
   820
                dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   821
                dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Isrel RS iffD1),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   822
                dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   823
                dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   824
                hyp_subst_tac, REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   825
            (rev (active_set_maps ~~ in_Isrels))])
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51761
diff changeset
   826
        (ctor_sets ~~ passive_set_maps),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   827
        rtac conjI,
49541
32fb6e4c7f4b renamed "map_simps" to "{c,d}tor_maps"
blanchet
parents: 49518
diff changeset
   828
        REPEAT_DETERM_N 2 o EVERY' [rtac trans, rtac ctor_map, rtac (ctor_inject RS iffD2),
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51739
diff changeset
   829
          rtac trans, rtac map_comp, rtac trans, rtac map_cong0,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   830
          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   831
          EVERY' (map (fn in_Isrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   832
            dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Isrel RS iffD1),
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
   833
            dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Isrels),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   834
          atac]]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   835
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   836
    EVERY' [rtac iffI, if_tac, only_if_tac] 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   837
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   838
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   839
end;