src/HOL/Cardinals/Cardinal_Order_Relation.thy
author wenzelm
Mon, 28 Dec 2015 17:43:30 +0100
changeset 61952 546958347e05
parent 61943 7fba644ed827
child 63040 eb4ddd18d635
permissions -rw-r--r--
prefer symbols for "Union", "Inter";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 48979
diff changeset
     1
(*  Title:      HOL/Cardinals/Cardinal_Order_Relation.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
Cardinal-order relations.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58127
diff changeset
     8
section {* Cardinal-Order Relations *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
theory Cardinal_Order_Relation
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 56191
diff changeset
    11
imports BNF_Cardinal_Order_Relation Wellorder_Constructions
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
declare
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
  card_order_on_well_order_on[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
  card_of_card_order_on[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  card_of_well_order_on[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
  Field_card_of[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
  card_of_Card_order[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
  card_of_Well_order[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
  card_of_least[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
  card_of_unique[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
  card_of_mono1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
  card_of_mono2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
  card_of_cong[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
  card_of_Field_ordLess[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
  card_of_Field_ordIso[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
  card_of_underS[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
  ordLess_Field[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
  card_of_empty[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
  card_of_empty1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
  card_of_image[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
  card_of_singl_ordLeq[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
  Card_order_singl_ordLeq[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
  card_of_Pow[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
  Card_order_Pow[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
  card_of_Plus1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
  Card_order_Plus1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
  card_of_Plus2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
  Card_order_Plus2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  card_of_Plus_mono1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
  card_of_Plus_mono2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
  card_of_Plus_mono[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
  card_of_Plus_cong2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
  card_of_Plus_cong[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
  card_of_Un_Plus_ordLeq[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
  card_of_Times1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
  card_of_Times2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
  card_of_Times3[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
  card_of_Times_mono1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
  card_of_Times_mono2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
  card_of_ordIso_finite[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
  card_of_Times_same_infinite[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
  card_of_Times_infinite_simps[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
  card_of_Plus_infinite1[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
  card_of_Plus_infinite2[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
  card_of_Plus_ordLess_infinite[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
  card_of_Plus_ordLess_infinite_Field[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
  infinite_cartesian_product[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
  cardSuc_Card_order[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
  cardSuc_greater[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
  cardSuc_ordLeq[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
  cardSuc_ordLeq_ordLess[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
  cardSuc_mono_ordLeq[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
  cardSuc_invar_ordIso[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
  card_of_cardSuc_finite[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
  cardSuc_finite[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
  card_of_Plus_ordLeq_infinite_Field[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
  curr_in[intro, simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
subsection {* Cardinal of a set *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
lemma card_of_inj_rel: assumes INJ: "!! x y y'. \<lbrakk>(x,y) : R; (x,y') : R\<rbrakk> \<Longrightarrow> y = y'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
shows "|{y. EX x. (x,y) : R}| <=o |{x. EX y. (x,y) : R}|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
  let ?Y = "{y. EX x. (x,y) : R}"  let ?X = "{x. EX y. (x,y) : R}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
  let ?f = "% y. SOME x. (x,y) : R"
51764
67f05cb13e08 optimized proofs
traytel
parents: 49310
diff changeset
    79
  have "?f ` ?Y <= ?X" by (auto dest: someI)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
  moreover have "inj_on ?f ?Y"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
  unfolding inj_on_def proof(auto)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
    fix y1 x1 y2 x2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
    assume *: "(x1, y1) \<in> R" "(x2, y2) \<in> R" and **: "?f y1 = ?f y2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
    hence "(?f y1,y1) : R" using someI[of "% x. (x,y1) : R"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
    moreover have "(?f y2,y2) : R" using * someI[of "% x. (x,y2) : R"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
    ultimately show "y1 = y2" using ** INJ by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
  ultimately show "|?Y| <=o |?X|" using card_of_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    89
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
lemma card_of_unique2: "\<lbrakk>card_order_on B r; bij_betw f A B\<rbrakk> \<Longrightarrow> r =o |A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
using card_of_ordIso card_of_unique ordIso_equivalence by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
lemma internalize_card_of_ordLess:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
"( |A| <o r) = (\<exists>B < Field r. |A| =o |B| \<and> |B| <o r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
  assume "|A| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
  then obtain p where 1: "Field p < Field r \<and> |A| =o p \<and> p <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
  using internalize_ordLess[of "|A|" r] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
  hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
  hence "|Field p| =o p" using card_of_Field_ordIso by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
  hence "|A| =o |Field p| \<and> |Field p| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
  using 1 ordIso_equivalence ordIso_ordLess_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
  thus "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r" using 1 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
  assume "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
  thus "|A| <o r" using ordIso_ordLess_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
lemma internalize_card_of_ordLess2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
"( |A| <o |C| ) = (\<exists>B < C. |A| =o |B| \<and> |B| <o |C| )"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
using internalize_card_of_ordLess[of "A" "|C|"] Field_card_of[of C] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
lemma Card_order_omax:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Card_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
shows "Card_order (omax R)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
  have "\<forall>r\<in>R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
  using assms unfolding card_order_on_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
  thus ?thesis using assms apply - apply(drule omax_in) by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
lemma Card_order_omax2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
assumes "finite I" and "I \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
shows "Card_order (omax {|A i| | i. i \<in> I})"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
  let ?R = "{|A i| | i. i \<in> I}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
  have "finite ?R" and "?R \<noteq> {}" using assms by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
  moreover have "\<forall>r\<in>?R. Card_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
  using card_of_Card_order by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
  ultimately show ?thesis by(rule Card_order_omax)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
subsection {* Cardinals versus set operations on arbitrary sets *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
54481
5c9819d7713b compile
blanchet
parents: 54475
diff changeset
   137
lemma card_of_set_type[simp]: "|UNIV::'a set| <o |UNIV::'a set set|"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   138
using card_of_Pow[of "UNIV::'a set"] by simp
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   139
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   140
lemma card_of_Un1[simp]:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   141
shows "|A| \<le>o |A \<union> B| "
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   142
using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   143
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   144
lemma card_of_diff[simp]:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   145
shows "|A - B| \<le>o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   146
using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   147
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   148
lemma subset_ordLeq_strict:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   149
assumes "A \<le> B" and "|A| <o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
shows "A < B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
  {assume "\<not>(A < B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   153
   hence "A = B" using assms(1) by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
   hence False using assms(2) not_ordLess_ordIso card_of_refl by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
  thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
corollary subset_ordLeq_diff:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
assumes "A \<le> B" and "|A| <o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   161
shows "B - A \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
using assms subset_ordLeq_strict by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
lemma card_of_empty4:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   165
"|{}::'b set| <o |A::'a set| = (A \<noteq> {})"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
proof(intro iffI notI)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
  assume *: "|{}::'b set| <o |A|" and "A = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
  hence "|A| =o |{}::'b set|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
  using card_of_ordIso unfolding bij_betw_def inj_on_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
  hence "|{}::'b set| =o |A|" using ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
  with * show False using not_ordLess_ordIso[of "|{}::'b set|" "|A|"] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
  assume "A \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
  hence "(\<not> (\<exists>f. inj_on f A \<and> f ` A \<subseteq> {}))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
  unfolding inj_on_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
  thus "| {} | <o | A |"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
  using card_of_ordLess by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
lemma card_of_empty5:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
"|A| <o |B| \<Longrightarrow> B \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   182
using card_of_empty not_ordLess_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
lemma Well_order_card_of_empty:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
"Well_order r \<Longrightarrow> |{}| \<le>o r" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
lemma card_of_UNIV[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
"|A :: 'a set| \<le>o |UNIV :: 'a set|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
using card_of_mono1[of A] by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
lemma card_of_UNIV2[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
"Card_order r \<Longrightarrow> (r :: 'a rel) \<le>o |UNIV :: 'a set|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
using card_of_UNIV[of "Field r"] card_of_Field_ordIso
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   194
      ordIso_symmetric ordIso_ordLeq_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   195
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
lemma card_of_Pow_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   197
assumes "|A| \<le>o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
shows "|Pow A| \<le>o |Pow B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
  obtain f where "inj_on f A \<and> f ` A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
  using assms card_of_ordLeq[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
  hence "inj_on (image f) (Pow A) \<and> (image f) ` (Pow A) \<le> (Pow B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   203
  by (auto simp add: inj_on_image_Pow image_Pow_mono)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
  thus ?thesis using card_of_ordLeq[of "Pow A"] by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
lemma ordIso_Pow_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
assumes "r \<le>o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
shows "|Pow(Field r)| \<le>o |Pow(Field r')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
using assms card_of_mono2 card_of_Pow_mono by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   211
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
lemma card_of_Pow_cong[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   213
assumes "|A| =o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   214
shows "|Pow A| =o |Pow B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   215
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
  obtain f where "bij_betw f A B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   217
  using assms card_of_ordIso[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   218
  hence "bij_betw (image f) (Pow A) (Pow B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
  by (auto simp add: bij_betw_image_Pow)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
  thus ?thesis using card_of_ordIso[of "Pow A"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   221
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   222
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   223
lemma ordIso_Pow_cong[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
assumes "r =o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   225
shows "|Pow(Field r)| =o |Pow(Field r')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   226
using assms card_of_cong card_of_Pow_cong by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
corollary Card_order_Plus_empty1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
"Card_order r \<Longrightarrow> r =o |(Field r) <+> {}|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
using card_of_Plus_empty1 card_of_Field_ordIso ordIso_equivalence by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
corollary Card_order_Plus_empty2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
"Card_order r \<Longrightarrow> r =o |{} <+> (Field r)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
using card_of_Plus_empty2 card_of_Field_ordIso ordIso_equivalence by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
lemma Card_order_Un1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
shows "Card_order r \<Longrightarrow> |Field r| \<le>o |(Field r) \<union> B| "
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   238
using card_of_Un1 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
lemma card_of_Un2[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
shows "|A| \<le>o |B \<union> A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
lemma Card_order_Un2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
shows "Card_order r \<Longrightarrow> |Field r| \<le>o |A \<union> (Field r)| "
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
using card_of_Un2 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   247
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
lemma Un_Plus_bij_betw:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
assumes "A Int B = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
shows "\<exists>f. bij_betw f (A \<union> B) (A <+> B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   252
  let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
  have "bij_betw ?f (A \<union> B) (A <+> B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   254
  using assms by(unfold bij_betw_def inj_on_def, auto)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
  thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
lemma card_of_Un_Plus_ordIso:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
assumes "A Int B = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   260
shows "|A \<union> B| =o |A <+> B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
using assms card_of_ordIso[of "A \<union> B"] Un_Plus_bij_betw[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
lemma card_of_Un_Plus_ordIso1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
"|A \<union> B| =o |A <+> (B - A)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
using card_of_Un_Plus_ordIso[of A "B - A"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
lemma card_of_Un_Plus_ordIso2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
"|A \<union> B| =o |(A - B) <+> B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
using card_of_Un_Plus_ordIso[of "A - B" B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
lemma card_of_Times_singl1: "|A| =o |A \<times> {b}|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
  have "bij_betw fst (A \<times> {b}) A" unfolding bij_betw_def inj_on_def by force
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   274
  thus ?thesis using card_of_ordIso ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   275
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   276
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
corollary Card_order_Times_singl1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
"Card_order r \<Longrightarrow> r =o |(Field r) \<times> {b}|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
using card_of_Times_singl1[of _ b] card_of_Field_ordIso ordIso_equivalence by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   280
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
lemma card_of_Times_singl2: "|A| =o |{b} \<times> A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
  have "bij_betw snd ({b} \<times> A) A" unfolding bij_betw_def inj_on_def by force
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
  thus ?thesis using card_of_ordIso ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
corollary Card_order_Times_singl2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
"Card_order r \<Longrightarrow> r =o |{a} \<times> (Field r)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
using card_of_Times_singl2[of _ a] card_of_Field_ordIso ordIso_equivalence by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
lemma card_of_Times_assoc: "|(A \<times> B) \<times> C| =o |A \<times> B \<times> C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
proof -
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
  let ?f = "\<lambda>((a,b),c). (a,(b,c))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
  have "A \<times> B \<times> C \<subseteq> ?f ` ((A \<times> B) \<times> C)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
  proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
    fix x assume "x \<in> A \<times> B \<times> C"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   297
    then obtain a b c where *: "a \<in> A" "b \<in> B" "c \<in> C" "x = (a, b, c)" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   298
    let ?x = "((a, b), c)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   299
    from * have "?x \<in> (A \<times> B) \<times> C" "x = ?f ?x" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
    thus "x \<in> ?f ` ((A \<times> B) \<times> C)" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   301
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   302
  hence "bij_betw ?f ((A \<times> B) \<times> C) (A \<times> B \<times> C)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
  unfolding bij_betw_def inj_on_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
  thus ?thesis using card_of_ordIso by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
qed
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
corollary Card_order_Times3:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   308
"Card_order r \<Longrightarrow> |Field r| \<le>o |(Field r) \<times> (Field r)|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   309
  by (rule card_of_Times3)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   310
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   311
lemma card_of_Times_cong1[simp]:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   312
assumes "|A| =o |B|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   313
shows "|A \<times> C| =o |B \<times> C|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   314
using assms by (simp add: ordIso_iff_ordLeq)
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   315
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   316
lemma card_of_Times_cong2[simp]:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   317
assumes "|A| =o |B|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   318
shows "|C \<times> A| =o |C \<times> B|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   319
using assms by (simp add: ordIso_iff_ordLeq)
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   320
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
lemma card_of_Times_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   322
assumes "|A| \<le>o |B|" and "|C| \<le>o |D|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
shows "|A \<times> C| \<le>o |B \<times> D|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
using assms card_of_Times_mono1[of A B C] card_of_Times_mono2[of C D B]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   325
      ordLeq_transitive[of "|A \<times> C|"] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
corollary ordLeq_Times_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   328
assumes "r \<le>o r'" and "p \<le>o p'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   329
shows "|(Field r) \<times> (Field p)| \<le>o |(Field r') \<times> (Field p')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Times_mono by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   331
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   332
corollary ordIso_Times_cong1[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   333
assumes "r =o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   334
shows "|(Field r) \<times> C| =o |(Field r') \<times> C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   335
using assms card_of_cong card_of_Times_cong1 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   336
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   337
corollary ordIso_Times_cong2:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   338
assumes "r =o r'"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   339
shows "|A \<times> (Field r)| =o |A \<times> (Field r')|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   340
using assms card_of_cong card_of_Times_cong2 by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   341
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   342
lemma card_of_Times_cong[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
assumes "|A| =o |B|" and "|C| =o |D|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
shows "|A \<times> C| =o |B \<times> D|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   345
using assms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   346
by (auto simp add: ordIso_iff_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   347
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
corollary ordIso_Times_cong:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   349
assumes "r =o r'" and "p =o p'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   350
shows "|(Field r) \<times> (Field p)| =o |(Field r') \<times> (Field p')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   351
using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Times_cong by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   352
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   353
lemma card_of_Sigma_mono2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   354
assumes "inj_on f (I::'i set)" and "f ` I \<le> (J::'j set)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   355
shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| \<le>o |SIGMA j : J. A j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   356
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
  let ?LEFT = "SIGMA i : I. A (f i)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   358
  let ?RIGHT = "SIGMA j : J. A j"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   359
  obtain u where u_def: "u = (\<lambda>(i::'i,a::'a). (f i,a))" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   360
  have "inj_on u ?LEFT \<and> u `?LEFT \<le> ?RIGHT"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
  using assms unfolding u_def inj_on_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   362
  thus ?thesis using card_of_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   363
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   364
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   365
lemma card_of_Sigma_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   366
assumes INJ: "inj_on f I" and IM: "f ` I \<le> J" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   367
        LEQ: "\<forall>j \<in> J. |A j| \<le>o |B j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   368
shows "|SIGMA i : I. A (f i)| \<le>o |SIGMA j : J. B j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   369
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   370
  have "\<forall>i \<in> I. |A(f i)| \<le>o |B(f i)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   371
  using IM LEQ by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   372
  hence "|SIGMA i : I. A (f i)| \<le>o |SIGMA i : I. B (f i)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   373
  using card_of_Sigma_mono1[of I] by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   374
  moreover have "|SIGMA i : I. B (f i)| \<le>o |SIGMA j : J. B j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   375
  using INJ IM card_of_Sigma_mono2 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   376
  ultimately show ?thesis using ordLeq_transitive by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   377
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   378
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
lemma ordLeq_Sigma_mono1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   380
assumes "\<forall>i \<in> I. p i \<le>o r i"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   381
shows "|SIGMA i : I. Field(p i)| \<le>o |SIGMA i : I. Field(r i)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
using assms by (auto simp add: card_of_Sigma_mono1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   384
lemma ordLeq_Sigma_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   385
assumes "inj_on f I" and "f ` I \<le> J" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   386
        "\<forall>j \<in> J. p j \<le>o r j"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   387
shows "|SIGMA i : I. Field(p(f i))| \<le>o |SIGMA j : J. Field(r j)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   388
using assms card_of_mono2 card_of_Sigma_mono
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   389
      [of f I J "\<lambda> i. Field(p i)" "\<lambda> j. Field(r j)"] by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   390
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   391
lemma card_of_Sigma_cong1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   392
assumes "\<forall>i \<in> I. |A i| =o |B i|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   393
shows "|SIGMA i : I. A i| =o |SIGMA i : I. B i|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   394
using assms by (auto simp add: card_of_Sigma_mono1 ordIso_iff_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   395
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   396
lemma card_of_Sigma_cong2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   397
assumes "bij_betw f (I::'i set) (J::'j set)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   398
shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| =o |SIGMA j : J. A j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   399
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   400
  let ?LEFT = "SIGMA i : I. A (f i)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   401
  let ?RIGHT = "SIGMA j : J. A j"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   402
  obtain u where u_def: "u = (\<lambda>(i::'i,a::'a). (f i,a))" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   403
  have "bij_betw u ?LEFT ?RIGHT"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   404
  using assms unfolding u_def bij_betw_def inj_on_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   405
  thus ?thesis using card_of_ordIso by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   406
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   407
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   408
lemma card_of_Sigma_cong:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
assumes BIJ: "bij_betw f I J" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
        ISO: "\<forall>j \<in> J. |A j| =o |B j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   411
shows "|SIGMA i : I. A (f i)| =o |SIGMA j : J. B j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   412
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   413
  have "\<forall>i \<in> I. |A(f i)| =o |B(f i)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   414
  using ISO BIJ unfolding bij_betw_def by blast
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   415
  hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|" by (rule card_of_Sigma_cong1)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   416
  moreover have "|SIGMA i : I. B (f i)| =o |SIGMA j : J. B j|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   417
  using BIJ card_of_Sigma_cong2 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
  ultimately show ?thesis using ordIso_transitive by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   419
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   420
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   421
lemma ordIso_Sigma_cong1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
assumes "\<forall>i \<in> I. p i =o r i"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   423
shows "|SIGMA i : I. Field(p i)| =o |SIGMA i : I. Field(r i)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   424
using assms by (auto simp add: card_of_Sigma_cong1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   425
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   426
lemma ordLeq_Sigma_cong:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   427
assumes "bij_betw f I J" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   428
        "\<forall>j \<in> J. p j =o r j"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   429
shows "|SIGMA i : I. Field(p(f i))| =o |SIGMA j : J. Field(r j)|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   430
using assms card_of_cong card_of_Sigma_cong
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   431
      [of f I J "\<lambda> j. Field(p j)" "\<lambda> j. Field(r j)"] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   433
lemma card_of_UNION_Sigma2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   434
assumes
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   435
"!! i j. \<lbrakk>{i,j} <= I; i ~= j\<rbrakk> \<Longrightarrow> A i Int A j = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   436
shows
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   437
"|\<Union>i\<in>I. A i| =o |Sigma I A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   438
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   439
  let ?L = "\<Union>i\<in>I. A i"  let ?R = "Sigma I A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   440
  have "|?L| <=o |?R|" using card_of_UNION_Sigma .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   441
  moreover have "|?R| <=o |?L|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   442
  proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   443
    have "inj_on snd ?R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   444
    unfolding inj_on_def using assms by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   445
    moreover have "snd ` ?R <= ?L" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   446
    ultimately show ?thesis using card_of_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   447
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   448
  ultimately show ?thesis by(simp add: ordIso_iff_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   449
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   450
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   451
corollary Plus_into_Times:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   452
assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   453
        B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   454
shows "\<exists>f. inj_on f (A <+> B) \<and> f ` (A <+> B) \<le> A \<times> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   455
using assms by (auto simp add: card_of_Plus_Times card_of_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   456
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   457
corollary Plus_into_Times_types:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   458
assumes A2: "(a1::'a) \<noteq> a2" and  B2: "(b1::'b) \<noteq> b2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   459
shows "\<exists>(f::'a + 'b \<Rightarrow> 'a * 'b). inj f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   460
using assms Plus_into_Times[of a1 a2 UNIV b1 b2 UNIV]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   461
by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   462
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   463
corollary Times_same_infinite_bij_betw:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   464
assumes "\<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   465
shows "\<exists>f. bij_betw f (A \<times> A) A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   466
using assms by (auto simp add: card_of_ordIso)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   467
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   468
corollary Times_same_infinite_bij_betw_types:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   469
assumes INF: "\<not>finite(UNIV::'a set)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   470
shows "\<exists>(f::('a * 'a) => 'a). bij f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   471
using assms Times_same_infinite_bij_betw[of "UNIV::'a set"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   472
by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   473
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   474
corollary Times_infinite_bij_betw:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   475
assumes INF: "\<not>finite A" and NE: "B \<noteq> {}" and INJ: "inj_on g B \<and> g ` B \<le> A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   476
shows "(\<exists>f. bij_betw f (A \<times> B) A) \<and> (\<exists>h. bij_betw h (B \<times> A) A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   477
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
  have "|B| \<le>o |A|" using INJ card_of_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   479
  thus ?thesis using INF NE
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   480
  by (auto simp add: card_of_ordIso card_of_Times_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   481
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   482
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   483
corollary Times_infinite_bij_betw_types:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   484
assumes INF: "\<not>finite(UNIV::'a set)" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   485
        BIJ: "inj(g::'b \<Rightarrow> 'a)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   486
shows "(\<exists>(f::('b * 'a) => 'a). bij f) \<and> (\<exists>(h::('a * 'b) => 'a). bij h)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   487
using assms Times_infinite_bij_betw[of "UNIV::'a set" "UNIV::'b set" g]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   488
by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   489
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   490
lemma card_of_Times_ordLeq_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   491
"\<lbrakk>\<not>finite C; |A| \<le>o |C|; |B| \<le>o |C|\<rbrakk>
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 60585
diff changeset
   492
 \<Longrightarrow> |A \<times> B| \<le>o |C|"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   493
by(simp add: card_of_Sigma_ordLeq_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   494
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   495
corollary Plus_infinite_bij_betw:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   496
assumes INF: "\<not>finite A" and INJ: "inj_on g B \<and> g ` B \<le> A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   497
shows "(\<exists>f. bij_betw f (A <+> B) A) \<and> (\<exists>h. bij_betw h (B <+> A) A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   498
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   499
  have "|B| \<le>o |A|" using INJ card_of_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   500
  thus ?thesis using INF
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   501
  by (auto simp add: card_of_ordIso)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   502
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   503
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   504
corollary Plus_infinite_bij_betw_types:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   505
assumes INF: "\<not>finite(UNIV::'a set)" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   506
        BIJ: "inj(g::'b \<Rightarrow> 'a)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   507
shows "(\<exists>(f::('b + 'a) => 'a). bij f) \<and> (\<exists>(h::('a + 'b) => 'a). bij h)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   508
using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   509
by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   510
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   511
lemma card_of_Un_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   512
assumes INF: "\<not>finite A" and LEQ: "|B| \<le>o |A|"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   513
shows "|A \<union> B| =o |A| \<and> |B \<union> A| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   514
proof-
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   515
  have "|A \<union> B| \<le>o |A <+> B|" by (rule card_of_Un_Plus_ordLeq)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   516
  moreover have "|A <+> B| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   517
  using assms by (metis card_of_Plus_infinite)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   518
  ultimately have "|A \<union> B| \<le>o |A|" using ordLeq_ordIso_trans by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   519
  hence "|A \<union> B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   520
  thus ?thesis using Un_commute[of B A] by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   521
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   522
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   523
lemma card_of_Un_infinite_simps[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   524
"\<lbrakk>\<not>finite A; |B| \<le>o |A| \<rbrakk> \<Longrightarrow> |A \<union> B| =o |A|"
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   525
"\<lbrakk>\<not>finite A; |B| \<le>o |A| \<rbrakk> \<Longrightarrow> |B \<union> A| =o |A|"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   526
using card_of_Un_infinite by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   527
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   528
lemma card_of_Un_diff_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   529
assumes INF: "\<not>finite A" and LESS: "|B| <o |A|"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   530
shows "|A - B| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   531
proof-
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   532
  obtain C where C_def: "C = A - B" by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   533
  have "|A \<union> B| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   534
  using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   535
  moreover have "C \<union> B = A \<union> B" unfolding C_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   536
  ultimately have 1: "|C \<union> B| =o |A|" by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   537
  (*  *)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   538
  {assume *: "|C| \<le>o |B|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   539
   moreover
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   540
   {assume **: "finite B"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   541
    hence "finite C"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   542
    using card_of_ordLeq_finite * by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   543
    hence False using ** INF card_of_ordIso_finite 1 by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   544
   }
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   545
   hence "\<not>finite B" by auto
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   546
   ultimately have False
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   547
   using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   548
  }
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   549
  hence 2: "|B| \<le>o |C|" using card_of_Well_order ordLeq_total by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   550
  {assume *: "finite C"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   551
    hence "finite B" using card_of_ordLeq_finite 2 by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   552
    hence False using * INF card_of_ordIso_finite 1 by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   553
  }
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   554
  hence "\<not>finite C" by auto
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   555
  hence "|C| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   556
  using  card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   557
  thus ?thesis unfolding C_def .
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   558
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   559
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   560
corollary Card_order_Un_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   561
assumes INF: "\<not>finite(Field r)" and CARD: "Card_order r" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   562
        LEQ: "p \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   563
shows "| (Field r) \<union> (Field p) | =o r \<and> | (Field p) \<union> (Field r) | =o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   564
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   565
  have "| Field r \<union> Field p | =o | Field r | \<and>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   566
        | Field p \<union> Field r | =o | Field r |"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   567
  using assms by (auto simp add: card_of_Un_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   568
  thus ?thesis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   569
  using assms card_of_Field_ordIso[of r]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   570
        ordIso_transitive[of "|Field r \<union> Field p|"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   571
        ordIso_transitive[of _ "|Field r|"] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   572
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   573
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   574
corollary subset_ordLeq_diff_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   575
assumes INF: "\<not>finite B" and SUB: "A \<le> B" and LESS: "|A| <o |B|"
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   576
shows "\<not>finite (B - A)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   577
using assms card_of_Un_diff_infinite card_of_ordIso_finite by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   578
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   579
lemma card_of_Times_ordLess_infinite[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   580
assumes INF: "\<not>finite C" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   581
        LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   582
shows "|A \<times> B| <o |C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   583
proof(cases "A = {} \<or> B = {}")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   584
  assume Case1: "A = {} \<or> B = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   585
  hence "A \<times> B = {}" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   586
  moreover have "C \<noteq> {}" using
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   587
  LESS1 card_of_empty5 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   588
  ultimately show ?thesis by(auto simp add:  card_of_empty4)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   589
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   590
  assume Case2: "\<not>(A = {} \<or> B = {})"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   591
  {assume *: "|C| \<le>o |A \<times> B|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   592
   hence "\<not>finite (A \<times> B)" using INF card_of_ordLeq_finite by blast
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   593
   hence 1: "\<not>finite A \<or> \<not>finite B" using finite_cartesian_product by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   594
   {assume Case21: "|A| \<le>o |B|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   595
    hence "\<not>finite B" using 1 card_of_ordLeq_finite by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   596
    hence "|A \<times> B| =o |B|" using Case2 Case21
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   597
    by (auto simp add: card_of_Times_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   598
    hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   599
   }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   600
   moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   601
   {assume Case22: "|B| \<le>o |A|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   602
    hence "\<not>finite A" using 1 card_of_ordLeq_finite by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   603
    hence "|A \<times> B| =o |A|" using Case2 Case22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   604
    by (auto simp add: card_of_Times_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   605
    hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   606
   }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   607
   ultimately have False using ordLeq_total card_of_Well_order[of A]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   608
   card_of_Well_order[of B] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   609
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   610
  thus ?thesis using ordLess_or_ordLeq[of "|A \<times> B|" "|C|"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   611
  card_of_Well_order[of "A \<times> B"] card_of_Well_order[of "C"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   612
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   613
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   614
lemma card_of_Times_ordLess_infinite_Field[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   615
assumes INF: "\<not>finite (Field r)" and r: "Card_order r" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   616
        LESS1: "|A| <o r" and LESS2: "|B| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   617
shows "|A \<times> B| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   618
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   619
  let ?C  = "Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   620
  have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   621
  ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   622
  hence "|A| <o |?C|"  "|B| <o |?C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   623
  using LESS1 LESS2 ordLess_ordIso_trans by blast+
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 60585
diff changeset
   624
  hence  "|A \<times> B| <o |?C|" using INF
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   625
  card_of_Times_ordLess_infinite by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   626
  thus ?thesis using 1 ordLess_ordIso_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   627
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   628
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   629
lemma card_of_Un_ordLess_infinite[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   630
assumes INF: "\<not>finite C" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   631
        LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   632
shows "|A \<union> B| <o |C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   633
using assms card_of_Plus_ordLess_infinite card_of_Un_Plus_ordLeq
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   634
      ordLeq_ordLess_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   635
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   636
lemma card_of_Un_ordLess_infinite_Field[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   637
assumes INF: "\<not>finite (Field r)" and r: "Card_order r" and
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   638
        LESS1: "|A| <o r" and LESS2: "|B| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   639
shows "|A Un B| <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   640
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   641
  let ?C  = "Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   642
  have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   643
  ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   644
  hence "|A| <o |?C|"  "|B| <o |?C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   645
  using LESS1 LESS2 ordLess_ordIso_trans by blast+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   646
  hence  "|A Un B| <o |?C|" using INF
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   647
  card_of_Un_ordLess_infinite by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   648
  thus ?thesis using 1 ordLess_ordIso_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   649
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   650
55174
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   651
lemma ordLeq_finite_Field:
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   652
assumes "r \<le>o s" and "finite (Field s)"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   653
shows "finite (Field r)"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   654
by (metis assms card_of_mono2 card_of_ordLeq_infinite)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   655
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   656
lemma ordIso_finite_Field:
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   657
assumes "r =o s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   658
shows "finite (Field r) \<longleftrightarrow> finite (Field s)"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   659
by (metis assms ordIso_iff_ordLeq ordLeq_finite_Field)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   660
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   661
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   662
subsection {* Cardinals versus set operations involving infinite sets *}
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   663
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   664
lemma finite_iff_cardOf_nat:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   665
"finite A = ( |A| <o |UNIV :: nat set| )"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   666
using infinite_iff_card_of_nat[of A]
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   667
not_ordLeq_iff_ordLess[of "|A|" "|UNIV :: nat set|"]
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   668
by fastforce
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   669
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   670
lemma finite_ordLess_infinite2[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   671
assumes "finite A" and "\<not>finite B"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   672
shows "|A| <o |B|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   673
using assms
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   674
finite_ordLess_infinite[of "|A|" "|B|"]
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   675
card_of_Well_order[of A] card_of_Well_order[of B]
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   676
Field_card_of[of A] Field_card_of[of B] by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   677
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   678
lemma infinite_card_of_insert:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   679
assumes "\<not>finite A"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   680
shows "|insert a A| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   681
proof-
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   682
  have iA: "insert a A = A \<union> {a}" by simp
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   683
  show ?thesis
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   684
  using infinite_imp_bij_betw2[OF assms] unfolding iA
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   685
  by (metis bij_betw_inv card_of_ordIso)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   686
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   687
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   688
lemma card_of_Un_singl_ordLess_infinite1:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   689
assumes "\<not>finite B" and "|A| <o |B|"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   690
shows "|{a} Un A| <o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   691
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   692
  have "|{a}| <o |B|" using assms by auto
51764
67f05cb13e08 optimized proofs
traytel
parents: 49310
diff changeset
   693
  thus ?thesis using assms card_of_Un_ordLess_infinite[of B] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   694
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   695
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   696
lemma card_of_Un_singl_ordLess_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   697
assumes "\<not>finite B"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   698
shows "( |A| <o |B| ) = ( |{a} Un A| <o |B| )"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   699
using assms card_of_Un_singl_ordLess_infinite1[of B A]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   700
proof(auto)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   701
  assume "|insert a A| <o |B|"
51764
67f05cb13e08 optimized proofs
traytel
parents: 49310
diff changeset
   702
  moreover have "|A| <=o |insert a A|" using card_of_mono1[of A "insert a A"] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   703
  ultimately show "|A| <o |B|" using ordLeq_ordLess_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   704
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   705
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   706
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   707
subsection {* Cardinals versus lists *}
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   708
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   709
text{* The next is an auxiliary operator, which shall be used for inductive
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   710
proofs of facts concerning the cardinality of @{text "List"} : *}
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   711
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   712
definition nlists :: "'a set \<Rightarrow> nat \<Rightarrow> 'a list set"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   713
where "nlists A n \<equiv> {l. set l \<le> A \<and> length l = n}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   714
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   715
lemma lists_def2: "lists A = {l. set l \<le> A}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   716
using in_listsI by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   717
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
   718
lemma lists_UNION_nlists: "lists A = (\<Union>n. nlists A n)"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   719
unfolding lists_def2 nlists_def by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   720
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   721
lemma card_of_lists: "|A| \<le>o |lists A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   722
proof-
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   723
  let ?h = "\<lambda> a. [a]"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   724
  have "inj_on ?h A \<and> ?h ` A \<le> lists A"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   725
  unfolding inj_on_def lists_def2 by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   726
  thus ?thesis by (metis card_of_ordLeq)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   727
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   728
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   729
lemma nlists_0: "nlists A 0 = {[]}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   730
unfolding nlists_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   731
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   732
lemma nlists_not_empty:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   733
assumes "A \<noteq> {}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   734
shows "nlists A n \<noteq> {}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   735
proof(induct n, simp add: nlists_0)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   736
  fix n assume "nlists A n \<noteq> {}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   737
  then obtain a and l where "a \<in> A \<and> l \<in> nlists A n" using assms by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   738
  hence "a # l \<in> nlists A (Suc n)" unfolding nlists_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   739
  thus "nlists A (Suc n) \<noteq> {}" by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   740
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   741
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   742
lemma Nil_in_lists: "[] \<in> lists A"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   743
unfolding lists_def2 by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   744
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   745
lemma lists_not_empty: "lists A \<noteq> {}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   746
using Nil_in_lists by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   747
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   748
lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   749
proof-
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   750
  let ?B = "A \<times> (nlists A n)"   let ?h = "\<lambda>(a,l). a # l"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   751
  have "inj_on ?h ?B \<and> ?h ` ?B \<le> nlists A (Suc n)"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   752
  unfolding inj_on_def nlists_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   753
  moreover have "nlists A (Suc n) \<le> ?h ` ?B"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   754
  proof(auto)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   755
    fix l assume "l \<in> nlists A (Suc n)"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   756
    hence 1: "length l = Suc n \<and> set l \<le> A" unfolding nlists_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   757
    then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   758
    hence "a \<in> A \<and> set l' \<le> A \<and> length l' = n" using 1 by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   759
    thus "l \<in> ?h ` ?B"  using 2 unfolding nlists_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   760
  qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   761
  ultimately have "bij_betw ?h ?B (nlists A (Suc n))"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   762
  unfolding bij_betw_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   763
  thus ?thesis using card_of_ordIso ordIso_symmetric by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   764
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   765
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   766
lemma card_of_nlists_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   767
assumes "\<not>finite A"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   768
shows "|nlists A n| \<le>o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   769
proof(induct n)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   770
  have "A \<noteq> {}" using assms by auto
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   771
  thus "|nlists A 0| \<le>o |A|" by (simp add: nlists_0)
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   772
next
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   773
  fix n assume IH: "|nlists A n| \<le>o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   774
  have "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   775
  using card_of_nlists_Succ by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   776
  moreover
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   777
  {have "nlists A n \<noteq> {}" using assms nlists_not_empty[of A] by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   778
   hence "|A \<times> (nlists A n)| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   779
   using assms IH by (auto simp add: card_of_Times_infinite)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   780
  }
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   781
  ultimately show "|nlists A (Suc n)| \<le>o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   782
  using ordIso_transitive ordIso_iff_ordLeq by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   783
qed
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   784
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   785
lemma Card_order_lists: "Card_order r \<Longrightarrow> r \<le>o |lists(Field r) |"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   786
using card_of_lists card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   787
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61943
diff changeset
   788
lemma Union_set_lists: "\<Union>(set ` (lists A)) = A"
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61943
diff changeset
   789
  unfolding lists_def2
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61943
diff changeset
   790
proof(auto)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   791
  fix a assume "a \<in> A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   792
  hence "set [a] \<le> A \<and> a \<in> set [a]" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   793
  thus "\<exists>l. set l \<le> A \<and> a \<in> set l" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   794
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   795
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   796
lemma inj_on_map_lists:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   797
assumes "inj_on f A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   798
shows "inj_on (map f) (lists A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   799
using assms Union_set_lists[of A] inj_on_mapI[of f "lists A"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   800
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   801
lemma map_lists_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   802
assumes "f ` A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   803
shows "(map f) ` (lists A) \<le> lists B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   804
using assms unfolding lists_def2 by (auto, blast) (* lethal combination of methods :)  *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   805
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   806
lemma map_lists_surjective:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   807
assumes "f ` A = B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   808
shows "(map f) ` (lists A) = lists B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   809
using assms unfolding lists_def2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   810
proof (auto, blast)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   811
  fix l' assume *: "set l' \<le> f ` A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   812
  have "set l' \<le> f ` A \<longrightarrow> l' \<in> map f ` {l. set l \<le> A}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   813
  proof(induct l', auto)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   814
    fix l a
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   815
    assume "a \<in> A" and "set l \<le> A" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   816
           IH: "f ` (set l) \<le> f ` A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   817
    hence "set (a # l) \<le> A" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   818
    hence "map f (a # l) \<in> map f ` {l. set l \<le> A}" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   819
    thus "f a # map f l \<in> map f ` {l. set l \<le> A}" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   820
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   821
  thus "l' \<in> map f ` {l. set l \<le> A}" using * by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   822
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   823
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   824
lemma bij_betw_map_lists:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   825
assumes "bij_betw f A B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   826
shows "bij_betw (map f) (lists A) (lists B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   827
using assms unfolding bij_betw_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   828
by(auto simp add: inj_on_map_lists map_lists_surjective)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   829
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   830
lemma card_of_lists_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   831
assumes "|A| \<le>o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   832
shows "|lists A| \<le>o |lists B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   833
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   834
  obtain f where "inj_on f A \<and> f ` A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   835
  using assms card_of_ordLeq[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   836
  hence "inj_on (map f) (lists A) \<and> (map f) ` (lists A) \<le> (lists B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   837
  by (auto simp add: inj_on_map_lists map_lists_mono)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   838
  thus ?thesis using card_of_ordLeq[of "lists A"] by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   839
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   840
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   841
lemma ordIso_lists_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   842
assumes "r \<le>o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   843
shows "|lists(Field r)| \<le>o |lists(Field r')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   844
using assms card_of_mono2 card_of_lists_mono by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   845
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   846
lemma card_of_lists_cong[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   847
assumes "|A| =o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   848
shows "|lists A| =o |lists B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   849
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   850
  obtain f where "bij_betw f A B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   851
  using assms card_of_ordIso[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   852
  hence "bij_betw (map f) (lists A) (lists B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   853
  by (auto simp add: bij_betw_map_lists)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   854
  thus ?thesis using card_of_ordIso[of "lists A"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   855
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   856
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   857
lemma card_of_lists_infinite[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   858
assumes "\<not>finite A"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   859
shows "|lists A| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   860
proof-
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   861
  have "|lists A| \<le>o |A|" unfolding lists_UNION_nlists
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   862
  by (rule card_of_UNION_ordLeq_infinite[OF assms _ ballI[OF card_of_nlists_infinite[OF assms]]])
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   863
    (metis infinite_iff_card_of_nat assms)
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   864
  thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   865
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   866
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   867
lemma Card_order_lists_infinite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   868
assumes "Card_order r" and "\<not>finite(Field r)"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   869
shows "|lists(Field r)| =o r"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   870
using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   871
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   872
lemma ordIso_lists_cong:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   873
assumes "r =o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   874
shows "|lists(Field r)| =o |lists(Field r')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   875
using assms card_of_cong card_of_lists_cong by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   876
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   877
corollary lists_infinite_bij_betw:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   878
assumes "\<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   879
shows "\<exists>f. bij_betw f (lists A) A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   880
using assms card_of_lists_infinite card_of_ordIso by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   881
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   882
corollary lists_infinite_bij_betw_types:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   883
assumes "\<not>finite(UNIV :: 'a set)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   884
shows "\<exists>(f::'a list \<Rightarrow> 'a). bij f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   885
using assms assms lists_infinite_bij_betw[of "UNIV::'a set"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   886
using lists_UNIV by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   887
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   888
55102
761e40ce91bc whitespace tuning
blanchet
parents: 55087
diff changeset
   889
subsection {* Cardinals versus the set-of-finite-sets operator *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   890
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   891
definition Fpow :: "'a set \<Rightarrow> 'a set set"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   892
where "Fpow A \<equiv> {X. X \<le> A \<and> finite X}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   893
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   894
lemma Fpow_mono: "A \<le> B \<Longrightarrow> Fpow A \<le> Fpow B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   895
unfolding Fpow_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   896
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   897
lemma empty_in_Fpow: "{} \<in> Fpow A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   898
unfolding Fpow_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   899
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   900
lemma Fpow_not_empty: "Fpow A \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   901
using empty_in_Fpow by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   902
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   903
lemma Fpow_subset_Pow: "Fpow A \<le> Pow A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   904
unfolding Fpow_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   905
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   906
lemma card_of_Fpow[simp]: "|A| \<le>o |Fpow A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   907
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   908
  let ?h = "\<lambda> a. {a}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   909
  have "inj_on ?h A \<and> ?h ` A \<le> Fpow A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   910
  unfolding inj_on_def Fpow_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   911
  thus ?thesis using card_of_ordLeq by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   912
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   913
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   914
lemma Card_order_Fpow: "Card_order r \<Longrightarrow> r \<le>o |Fpow(Field r) |"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   915
using card_of_Fpow card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   916
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   917
lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   918
unfolding Fpow_def Pow_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   919
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   920
lemma inj_on_image_Fpow:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   921
assumes "inj_on f A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   922
shows "inj_on (image f) (Fpow A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   923
using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   924
      inj_on_image_Pow by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   925
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   926
lemma image_Fpow_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   927
assumes "f ` A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   928
shows "(image f) ` (Fpow A) \<le> Fpow B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   929
using assms by(unfold Fpow_def, auto)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   930
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   931
lemma image_Fpow_surjective:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   932
assumes "f ` A = B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   933
shows "(image f) ` (Fpow A) = Fpow B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   934
using assms proof(unfold Fpow_def, auto)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   935
  fix Y assume *: "Y \<le> f ` A" and **: "finite Y"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   936
  hence "\<forall>b \<in> Y. \<exists>a. a \<in> A \<and> f a = b" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   937
  with bchoice[of Y "\<lambda>b a. a \<in> A \<and> f a = b"]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   938
  obtain g where 1: "\<forall>b \<in> Y. g b \<in> A \<and> f(g b) = b" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   939
  obtain X where X_def: "X = g ` Y" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   940
  have "f ` X = Y \<and> X \<le> A \<and> finite X"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   941
  by(unfold X_def, force simp add: ** 1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   942
  thus "Y \<in> (image f) ` {X. X \<le> A \<and> finite X}" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   943
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   944
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   945
lemma bij_betw_image_Fpow:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   946
assumes "bij_betw f A B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   947
shows "bij_betw (image f) (Fpow A) (Fpow B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   948
using assms unfolding bij_betw_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   949
by (auto simp add: inj_on_image_Fpow image_Fpow_surjective)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   950
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   951
lemma card_of_Fpow_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   952
assumes "|A| \<le>o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   953
shows "|Fpow A| \<le>o |Fpow B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   954
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   955
  obtain f where "inj_on f A \<and> f ` A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   956
  using assms card_of_ordLeq[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   957
  hence "inj_on (image f) (Fpow A) \<and> (image f) ` (Fpow A) \<le> (Fpow B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   958
  by (auto simp add: inj_on_image_Fpow image_Fpow_mono)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   959
  thus ?thesis using card_of_ordLeq[of "Fpow A"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   960
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   961
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   962
lemma ordIso_Fpow_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   963
assumes "r \<le>o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   964
shows "|Fpow(Field r)| \<le>o |Fpow(Field r')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   965
using assms card_of_mono2 card_of_Fpow_mono by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   966
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   967
lemma card_of_Fpow_cong[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   968
assumes "|A| =o |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   969
shows "|Fpow A| =o |Fpow B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   970
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   971
  obtain f where "bij_betw f A B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   972
  using assms card_of_ordIso[of A B] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   973
  hence "bij_betw (image f) (Fpow A) (Fpow B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   974
  by (auto simp add: bij_betw_image_Fpow)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   975
  thus ?thesis using card_of_ordIso[of "Fpow A"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   976
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   977
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   978
lemma ordIso_Fpow_cong:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   979
assumes "r =o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   980
shows "|Fpow(Field r)| =o |Fpow(Field r')|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   981
using assms card_of_cong card_of_Fpow_cong by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   982
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   983
lemma card_of_Fpow_lists: "|Fpow A| \<le>o |lists A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   984
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   985
  have "set ` (lists A) = Fpow A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   986
  unfolding lists_def2 Fpow_def using finite_list finite_set by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   987
  thus ?thesis using card_of_ordLeq2[of "Fpow A"] Fpow_not_empty[of A] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   988
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   989
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   990
lemma card_of_Fpow_infinite[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   991
assumes "\<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   992
shows "|Fpow A| =o |A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   993
using assms card_of_Fpow_lists card_of_lists_infinite card_of_Fpow
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   994
      ordLeq_ordIso_trans ordIso_iff_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   995
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   996
corollary Fpow_infinite_bij_betw:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
   997
assumes "\<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   998
shows "\<exists>f. bij_betw f (Fpow A) A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   999
using assms card_of_Fpow_infinite card_of_ordIso by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1000
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1001
55102
761e40ce91bc whitespace tuning
blanchet
parents: 55087
diff changeset
  1002
subsection {* The cardinal $\omega$ and the finite cardinals *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1003
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1004
subsubsection {* First as well-orders *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1005
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1006
lemma Field_natLess: "Field natLess = (UNIV::nat set)"
56011
39d5043ce8a3 made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents: 55174
diff changeset
  1007
by(unfold Field_def natLess_def, auto)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1008
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1009
lemma natLeq_well_order_on: "well_order_on UNIV natLeq"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1010
using natLeq_Well_order Field_natLeq by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1011
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1012
lemma natLeq_wo_rel: "wo_rel natLeq"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1013
unfolding wo_rel_def using natLeq_Well_order .
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1014
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1015
lemma natLeq_ofilter_less: "ofilter natLeq {0 ..< n}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1016
by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def,
56011
39d5043ce8a3 made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents: 55174
diff changeset
  1017
   simp add: Field_natLeq, unfold under_def natLeq_def, auto)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1018
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1019
lemma natLeq_ofilter_leq: "ofilter natLeq {0 .. n}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1020
by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def,
56011
39d5043ce8a3 made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents: 55174
diff changeset
  1021
   simp add: Field_natLeq, unfold under_def natLeq_def, auto)
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1022
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1023
lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1024
using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1025
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1026
lemma closed_nat_set_iff:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1027
assumes "\<forall>(m::nat) n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1028
shows "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1029
proof-
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1030
  {assume "A \<noteq> UNIV" hence "\<exists>n. n \<notin> A" by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1031
   moreover obtain n where n_def: "n = (LEAST n. n \<notin> A)" by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1032
   ultimately have 1: "n \<notin> A \<and> (\<forall>m. m < n \<longrightarrow> m \<in> A)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1033
   using LeastI_ex[of "\<lambda> n. n \<notin> A"] n_def Least_le[of "\<lambda> n. n \<notin> A"] by fastforce
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1034
   have "A = {0 ..< n}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1035
   proof(auto simp add: 1)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1036
     fix m assume *: "m \<in> A"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1037
     {assume "n \<le> m" with assms * have "n \<in> A" by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1038
      hence False using 1 by auto
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1039
     }
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1040
     thus "m < n" by fastforce
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1041
   qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1042
   hence "\<exists>n. A = {0 ..< n}" by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1043
  }
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1044
  thus ?thesis by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1045
qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1046
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1047
lemma natLeq_ofilter_iff:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1048
"ofilter natLeq A = (A = UNIV \<or> (\<exists>n. A = {0 ..< n}))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1049
proof(rule iffI)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1050
  assume "ofilter natLeq A"
56011
39d5043ce8a3 made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents: 55174
diff changeset
  1051
  hence "\<forall>m n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A" using natLeq_wo_rel
39d5043ce8a3 made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents: 55174
diff changeset
  1052
  by(auto simp add: natLeq_def wo_rel.ofilter_def under_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1053
  thus "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1054
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1055
  assume "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1056
  thus "ofilter natLeq A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1057
  by(auto simp add: natLeq_ofilter_less natLeq_UNIV_ofilter)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1058
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1059
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1060
lemma natLeq_under_leq: "under natLeq n = {0 .. n}"
56011
39d5043ce8a3 made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents: 55174
diff changeset
  1061
unfolding under_def natLeq_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1062
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1063
lemma natLeq_on_ofilter_less_eq:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1064
"n \<le> m \<Longrightarrow> wo_rel.ofilter (natLeq_on m) {0 ..< n}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1065
apply (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1066
apply (simp add: Field_natLeq_on)
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1067
by (auto simp add: under_def)
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1068
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1069
lemma natLeq_on_ofilter_iff:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1070
"wo_rel.ofilter (natLeq_on m) A = (\<exists>n \<le> m. A = {0 ..< n})"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1071
proof(rule iffI)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1072
  assume *: "wo_rel.ofilter (natLeq_on m) A"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1073
  hence 1: "A \<le> {0..<m}"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1074
  by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def under_def Field_natLeq_on)
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1075
  hence "\<forall>n1 n2. n2 \<in> A \<and> n1 \<le> n2 \<longrightarrow> n1 \<in> A"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1076
  using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def under_def)
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1077
  hence "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1078
  thus "\<exists>n \<le> m. A = {0 ..< n}" using 1 atLeastLessThan_less_eq by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1079
next
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1080
  assume "(\<exists>n\<le>m. A = {0 ..< n})"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1081
  thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp add: natLeq_on_ofilter_less_eq)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1082
qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1083
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1084
corollary natLeq_on_ofilter:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1085
"ofilter(natLeq_on n) {0 ..< n}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1086
by (auto simp add: natLeq_on_ofilter_less_eq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1087
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1088
lemma natLeq_on_ofilter_less:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1089
"n < m \<Longrightarrow> ofilter (natLeq_on m) {0 .. n}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1090
by(auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def,
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1091
   simp add: Field_natLeq_on, unfold under_def, auto)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1092
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1093
lemma natLeq_on_ordLess_natLeq: "natLeq_on n <o natLeq"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1094
using Field_natLeq Field_natLeq_on[of n]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1095
      finite_ordLess_infinite[of "natLeq_on n" natLeq]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1096
      natLeq_Well_order natLeq_on_Well_order[of n] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1097
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1098
lemma natLeq_on_injective:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1099
"natLeq_on m = natLeq_on n \<Longrightarrow> m = n"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1100
using Field_natLeq_on[of m] Field_natLeq_on[of n]
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1101
      atLeastLessThan_injective[of m n, unfolded atLeastLessThan_def] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1102
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1103
lemma natLeq_on_injective_ordIso:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1104
"(natLeq_on m =o natLeq_on n) = (m = n)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1105
proof(auto simp add: natLeq_on_Well_order ordIso_reflexive)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1106
  assume "natLeq_on m =o natLeq_on n"
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1107
  then obtain f where "bij_betw f {x. x<m} {x. x<n}"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1108
  using Field_natLeq_on assms unfolding ordIso_def iso_def[abs_def] by auto
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1109
  thus "m = n" using atLeastLessThan_injective2[of f m n]
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1110
    unfolding atLeast_0 atLeastLessThan_def lessThan_def Int_UNIV_left by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1111
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1112
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1113
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1114
subsubsection {* Then as cardinals *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1115
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1116
lemma ordIso_natLeq_infinite1:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1117
"|A| =o natLeq \<Longrightarrow> \<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1118
using ordIso_symmetric ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1119
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1120
lemma ordIso_natLeq_infinite2:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1121
"natLeq =o |A| \<Longrightarrow> \<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1122
using ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1123
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1124
lemma ordIso_natLeq_on_imp_finite:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1125
"|A| =o natLeq_on n \<Longrightarrow> finite A"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1126
unfolding ordIso_def iso_def[abs_def]
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1127
by (auto simp: Field_natLeq_on bij_betw_finite)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1128
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1129
lemma natLeq_on_Card_order: "Card_order (natLeq_on n)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1130
proof(unfold card_order_on_def,
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1131
      auto simp add: natLeq_on_Well_order, simp add: Field_natLeq_on)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1132
  fix r assume "well_order_on {x. x < n} r"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1133
  thus "natLeq_on n \<le>o r"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1134
  using finite_atLeastLessThan natLeq_on_well_order_on
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1135
        finite_well_order_on_ordIso ordIso_iff_ordLeq by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1136
qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1137
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1138
corollary card_of_Field_natLeq_on:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1139
"|Field (natLeq_on n)| =o natLeq_on n"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1140
using Field_natLeq_on natLeq_on_Card_order
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1141
      Card_order_iff_ordIso_card_of[of "natLeq_on n"]
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1142
      ordIso_symmetric[of "natLeq_on n"] by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1143
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1144
corollary card_of_less:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1145
"|{0 ..< n}| =o natLeq_on n"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1146
using Field_natLeq_on card_of_Field_natLeq_on
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1147
unfolding atLeast_0 atLeastLessThan_def lessThan_def Int_UNIV_left by auto
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1148
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1149
lemma natLeq_on_ordLeq_less_eq:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1150
"((natLeq_on m) \<le>o (natLeq_on n)) = (m \<le> n)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1151
proof
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1152
  assume "natLeq_on m \<le>o natLeq_on n"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1153
  then obtain f where "inj_on f {x. x < m} \<and> f ` {x. x < m} \<le> {x. x < n}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1154
  unfolding ordLeq_def using
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1155
    embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on]
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1156
     embed_Field[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1157
  thus "m \<le> n" using atLeastLessThan_less_eq2
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1158
    unfolding atLeast_0 atLeastLessThan_def lessThan_def Int_UNIV_left by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1159
next
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1160
  assume "m \<le> n"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1161
  hence "inj_on id {0..<m} \<and> id ` {0..<m} \<le> {0..<n}" unfolding inj_on_def by auto
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1162
  hence "|{0..<m}| \<le>o |{0..<n}|" using card_of_ordLeq by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1163
  thus "natLeq_on m \<le>o natLeq_on n"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1164
  using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1165
qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1166
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1167
lemma natLeq_on_ordLeq_less:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1168
"((natLeq_on m) <o (natLeq_on n)) = (m < n)"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1169
using not_ordLeq_iff_ordLess[OF natLeq_on_Well_order natLeq_on_Well_order, of n m]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1170
   natLeq_on_ordLeq_less_eq[of n m] by linarith
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1171
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1172
lemma ordLeq_natLeq_on_imp_finite:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1173
assumes "|A| \<le>o natLeq_on n"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1174
shows "finite A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1175
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1176
  have "|A| \<le>o |{0 ..< n}|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1177
  using assms card_of_less ordIso_symmetric ordLeq_ordIso_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1178
  thus ?thesis by (auto simp add: card_of_ordLeq_finite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1179
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1180
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1181
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1182
subsubsection {* "Backward compatibility" with the numeric cardinal operator for finite sets *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1183
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1184
lemma finite_card_of_iff_card2:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1185
assumes FIN: "finite A" and FIN': "finite B"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1186
shows "( |A| \<le>o |B| ) = (card A \<le> card B)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1187
using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1188
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1189
lemma finite_imp_card_of_natLeq_on:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1190
assumes "finite A"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1191
shows "|A| =o natLeq_on (card A)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1192
proof-
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1193
  obtain h where "bij_betw h A {0 ..< card A}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1194
  using assms ex_bij_betw_finite_nat by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1195
  thus ?thesis using card_of_ordIso card_of_less ordIso_equivalence by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1196
qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1197
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1198
lemma finite_iff_card_of_natLeq_on:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1199
"finite A = (\<exists>n. |A| =o natLeq_on n)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1200
using finite_imp_card_of_natLeq_on[of A]
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1201
by(auto simp add: ordIso_natLeq_on_imp_finite)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1202
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1203
lemma finite_card_of_iff_card:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1204
assumes FIN: "finite A" and FIN': "finite B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1205
shows "( |A| =o |B| ) = (card A = card B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1206
using assms card_of_ordIso[of A B] bij_betw_iff_card[of A B] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1207
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1208
lemma finite_card_of_iff_card3:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1209
assumes FIN: "finite A" and FIN': "finite B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1210
shows "( |A| <o |B| ) = (card A < card B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1211
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1212
  have "( |A| <o |B| ) = (~ ( |B| \<le>o |A| ))" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1213
  also have "... = (~ (card B \<le> card A))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1214
  using assms by(simp add: finite_card_of_iff_card2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1215
  also have "... = (card A < card B)" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1216
  finally show ?thesis .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1217
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1218
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1219
lemma card_Field_natLeq_on:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1220
"card(Field(natLeq_on n)) = n"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1221
using Field_natLeq_on card_atLeastLessThan by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1222
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1223
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1224
subsection {* The successor of a cardinal *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1225
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1226
lemma embed_implies_ordIso_Restr:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1227
assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1228
shows "r' =o Restr r (f ` (Field r'))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1229
using assms embed_implies_iso_Restr Well_order_Restr unfolding ordIso_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1230
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1231
lemma cardSuc_Well_order[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1232
"Card_order r \<Longrightarrow> Well_order(cardSuc r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1233
using cardSuc_Card_order unfolding card_order_on_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1234
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1235
lemma Field_cardSuc_not_empty:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1236
assumes "Card_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1237
shows "Field (cardSuc r) \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1238
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1239
  assume "Field(cardSuc r) = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1240
  hence "|Field(cardSuc r)| \<le>o r" using assms Card_order_empty[of r] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1241
  hence "cardSuc r \<le>o r" using assms card_of_Field_ordIso
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1242
  cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1243
  thus False using cardSuc_greater not_ordLess_ordLeq assms by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1244
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1245
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1246
lemma cardSuc_mono_ordLess[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1247
assumes CARD: "Card_order r" and CARD': "Card_order r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1248
shows "(cardSuc r <o cardSuc r') = (r <o r')"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1249
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1250
  have 0: "Well_order r \<and> Well_order r' \<and> Well_order(cardSuc r) \<and> Well_order(cardSuc r')"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1251
  using assms by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1252
  thus ?thesis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1253
  using not_ordLeq_iff_ordLess not_ordLeq_iff_ordLess[of r r']
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1254
  using cardSuc_mono_ordLeq[of r' r] assms by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1255
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1256
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1257
lemma cardSuc_natLeq_on_Suc:
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1258
"cardSuc(natLeq_on n) =o natLeq_on(Suc n)"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1259
proof-
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1260
  obtain r r' p where r_def: "r = natLeq_on n" and
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1261
                      r'_def: "r' = cardSuc(natLeq_on n)"  and
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1262
                      p_def: "p = natLeq_on(Suc n)" by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1263
  (* Preliminary facts:  *)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1264
  have CARD: "Card_order r \<and> Card_order r' \<and> Card_order p" unfolding r_def r'_def p_def
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1265
  using cardSuc_ordLess_ordLeq natLeq_on_Card_order cardSuc_Card_order by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1266
  hence WELL: "Well_order r \<and> Well_order r' \<and>  Well_order p"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1267
  unfolding card_order_on_def by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1268
  have FIELD: "Field r = {0..<n} \<and> Field p = {0..<(Suc n)}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1269
  unfolding r_def p_def Field_natLeq_on atLeast_0 atLeastLessThan_def lessThan_def by simp
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1270
  hence FIN: "finite (Field r)" by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1271
  have "r <o r'" using CARD unfolding r_def r'_def using cardSuc_greater by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1272
  hence "|Field r| <o r'" using CARD card_of_Field_ordIso ordIso_ordLess_trans by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1273
  hence LESS: "|Field r| <o |Field r'|"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1274
  using CARD card_of_Field_ordIso ordLess_ordIso_trans ordIso_symmetric by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1275
  (* Main proof: *)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1276
  have "r' \<le>o p" using CARD unfolding r_def r'_def p_def
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1277
  using natLeq_on_ordLeq_less cardSuc_ordLess_ordLeq by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1278
  moreover have "p \<le>o r'"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1279
  proof-
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1280
    {assume "r' <o p"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1281
     then obtain f where 0: "embedS r' p f" unfolding ordLess_def by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1282
     let ?q = "Restr p (f ` Field r')"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1283
     have 1: "embed r' p f" using 0 unfolding embedS_def by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1284
     hence 2: "f ` Field r' < {0..<(Suc n)}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1285
     using WELL FIELD 0 by (auto simp add: embedS_iff)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1286
     have "wo_rel.ofilter p (f ` Field r')" using embed_Field_ofilter 1 WELL by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1287
     then obtain m where "m \<le> Suc n" and 3: "f ` (Field r') = {0..<m}"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1288
     unfolding p_def by (auto simp add: natLeq_on_ofilter_iff)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1289
     hence 4: "m \<le> n" using 2 by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1290
     (*  *)
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1291
     have "bij_betw f (Field r') (f ` (Field r'))"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1292
     using WELL embed_inj_on[OF _ 1] unfolding bij_betw_def by blast
54581
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1293
     moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1294
     ultimately have 5: "finite (Field r') \<and> card(Field r') = card (f ` (Field r'))"
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1295
     using bij_betw_same_card bij_betw_finite by metis
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1296
     hence "card(Field r') \<le> card(Field r)" using 3 4 FIELD by force
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1297
     hence "|Field r'| \<le>o |Field r|" using FIN 5 finite_card_of_iff_card2 by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1298
     hence False using LESS not_ordLess_ordLeq by auto
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1299
    }
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1300
    thus ?thesis using WELL CARD by fastforce
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1301
  qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1302
  ultimately show ?thesis using ordIso_iff_ordLeq unfolding r'_def p_def by blast
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1303
qed
1502a1f707d9 eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents: 54578
diff changeset
  1304
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1305
lemma card_of_Plus_ordLeq_infinite[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1306
assumes C: "\<not>finite C" and A: "|A| \<le>o |C|" and B: "|B| \<le>o |C|"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1307
shows "|A <+> B| \<le>o |C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1308
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1309
  let ?r = "cardSuc |C|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1310
  have "Card_order ?r \<and> \<not>finite (Field ?r)" using assms by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1311
  moreover have "|A| <o ?r" and "|B| <o ?r" using A B by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1312
  ultimately have "|A <+> B| <o ?r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1313
  using card_of_Plus_ordLess_infinite_Field by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1314
  thus ?thesis using C by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1315
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1316
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1317
lemma card_of_Un_ordLeq_infinite[simp]:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1318
assumes C: "\<not>finite C" and A: "|A| \<le>o |C|" and B: "|B| \<le>o |C|"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1319
shows "|A Un B| \<le>o |C|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1320
using assms card_of_Plus_ordLeq_infinite card_of_Un_Plus_ordLeq
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1321
ordLeq_transitive by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1322
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1323
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1324
subsection {* Others *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1325
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1326
lemma under_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1327
assumes "Well_order r" and "(i,j) \<in> r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1328
shows "under r i \<subseteq> under r j"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1329
using assms unfolding under_def order_on_defs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1330
trans_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1331
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1332
lemma underS_under:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1333
assumes "i \<in> Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1334
shows "underS r i = under r i - {i}"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1335
using assms unfolding underS_def under_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1336
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1337
lemma relChain_under:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1338
assumes "Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1339
shows "relChain r (\<lambda> i. under r i)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1340
using assms unfolding relChain_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1341
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1342
lemma card_of_infinite_diff_finite:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1343
assumes "\<not>finite A" and "finite B"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1344
shows "|A - B| =o |A|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1345
by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1346
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1347
lemma infinite_card_of_diff_singl:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1348
assumes "\<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1349
shows "|A - {a}| =o |A|"
52544
0c4b140cff00 tuned spelling
traytel
parents: 51764
diff changeset
  1350
by (metis assms card_of_infinite_diff_finite finite.emptyI finite_insert)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1351
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1352
lemma card_of_vimage:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1353
assumes "B \<subseteq> range f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1354
shows "|B| \<le>o |f -` B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1355
apply(rule surj_imp_ordLeq[of _ f])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1356
using assms by (metis Int_absorb2 image_vimage_eq order_refl)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1357
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1358
lemma surj_card_of_vimage:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1359
assumes "surj f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1360
shows "|B| \<le>o |f -` B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1361
by (metis assms card_of_vimage subset_UNIV)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1362
54794
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1363
lemma infinite_Pow:
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1364
assumes "\<not> finite A"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1365
shows "\<not> finite (Pow A)"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1366
proof-
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1367
  have "|A| \<le>o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1368
  thus ?thesis by (metis assms finite_Pow_iff)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1369
qed
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
  1370
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1371
(* bounded powerset *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1372
definition Bpow where
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1373
"Bpow r A \<equiv> {X . X \<subseteq> A \<and> |X| \<le>o r}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1374
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1375
lemma Bpow_empty[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1376
assumes "Card_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1377
shows "Bpow r {} = {{}}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1378
using assms unfolding Bpow_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1379
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1380
lemma singl_in_Bpow:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1381
assumes rc: "Card_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1382
and r: "Field r \<noteq> {}" and a: "a \<in> A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1383
shows "{a} \<in> Bpow r A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1384
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1385
  have "|{a}| \<le>o r" using r rc by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1386
  thus ?thesis unfolding Bpow_def using a by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1387
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1388
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1389
lemma ordLeq_card_Bpow:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1390
assumes rc: "Card_order r" and r: "Field r \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1391
shows "|A| \<le>o |Bpow r A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1392
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1393
  have "inj_on (\<lambda> a. {a}) A" unfolding inj_on_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1394
  moreover have "(\<lambda> a. {a}) ` A \<subseteq> Bpow r A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1395
  using singl_in_Bpow[OF assms] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1396
  ultimately show ?thesis unfolding card_of_ordLeq[symmetric] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1397
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1398
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1399
lemma infinite_Bpow:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1400
assumes rc: "Card_order r" and r: "Field r \<noteq> {}"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1401
and A: "\<not>finite A"
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1402
shows "\<not>finite (Bpow r A)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1403
using ordLeq_card_Bpow[OF rc r]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1404
by (metis A card_of_ordLeq_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1405
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1406
definition Func_option where
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1407
 "Func_option A B \<equiv>
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1408
  {f. (\<forall> a. f a \<noteq> None \<longleftrightarrow> a \<in> A) \<and> (\<forall> a \<in> A. case f a of Some b \<Rightarrow> b \<in> B |None \<Rightarrow> True)}"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1409
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1410
lemma card_of_Func_option_Func:
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1411
"|Func_option A B| =o |Func A B|"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1412
proof (rule ordIso_symmetric, unfold card_of_ordIso[symmetric], intro exI)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1413
  let ?F = "\<lambda> f a. if a \<in> A then Some (f a) else None"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1414
  show "bij_betw ?F (Func A B) (Func_option A B)"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1415
  unfolding bij_betw_def unfolding inj_on_def proof(intro conjI ballI impI)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1416
    fix f g assume f: "f \<in> Func A B" and g: "g \<in> Func A B" and eq: "?F f = ?F g"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1417
    show "f = g"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1418
    proof(rule ext)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1419
      fix a
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1420
      show "f a = g a"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1421
      proof(cases "a \<in> A")
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1422
        case True
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1423
        have "Some (f a) = ?F f a" using True by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1424
        also have "... = ?F g a" using eq unfolding fun_eq_iff by(rule allE)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1425
        also have "... = Some (g a)" using True by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1426
        finally have "Some (f a) = Some (g a)" .
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1427
        thus ?thesis by simp
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1428
      qed(insert f g, unfold Func_def, auto)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1429
    qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1430
  next
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1431
    show "?F ` Func A B = Func_option A B"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1432
    proof safe
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1433
      fix f assume f: "f \<in> Func_option A B"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1434
      def g \<equiv> "\<lambda> a. case f a of Some b \<Rightarrow> b | None \<Rightarrow> undefined"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1435
      have "g \<in> Func A B"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1436
      using f unfolding g_def Func_def Func_option_def by force+
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1437
      moreover have "f = ?F g"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1438
      proof(rule ext)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1439
        fix a show "f a = ?F g a"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1440
        using f unfolding Func_option_def g_def by (cases "a \<in> A") force+
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1441
      qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1442
      ultimately show "f \<in> ?F ` (Func A B)" by blast
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1443
    qed(unfold Func_def Func_option_def, auto)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1444
  qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1445
qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1446
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1447
(* partial-function space: *)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1448
definition Pfunc where
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1449
"Pfunc A B \<equiv>
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1450
 {f. (\<forall>a. f a \<noteq> None \<longrightarrow> a \<in> A) \<and>
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1451
     (\<forall>a. case f a of None \<Rightarrow> True | Some b \<Rightarrow> b \<in> B)}"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1452
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1453
lemma Func_Pfunc:
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1454
"Func_option A B \<subseteq> Pfunc A B"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1455
unfolding Func_option_def Pfunc_def by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1456
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1457
lemma Pfunc_Func_option:
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1458
"Pfunc A B = (\<Union>A' \<in> Pow A. Func_option A' B)"
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1459
proof safe
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1460
  fix f assume f: "f \<in> Pfunc A B"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1461
  show "f \<in> (\<Union>A'\<in>Pow A. Func_option A' B)"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1462
  proof (intro UN_I)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1463
    let ?A' = "{a. f a \<noteq> None}"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1464
    show "?A' \<in> Pow A" using f unfolding Pow_def Pfunc_def by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1465
    show "f \<in> Func_option ?A' B" using f unfolding Func_option_def Pfunc_def by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1466
  qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1467
next
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1468
  fix f A' assume "f \<in> Func_option A' B" and "A' \<subseteq> A"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1469
  thus "f \<in> Pfunc A B" unfolding Func_option_def Pfunc_def by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1470
qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1471
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1472
lemma card_of_Func_mono:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1473
fixes A1 A2 :: "'a set" and B :: "'b set"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1474
assumes A12: "A1 \<subseteq> A2" and B: "B \<noteq> {}"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1475
shows "|Func A1 B| \<le>o |Func A2 B|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1476
proof-
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1477
  obtain bb where bb: "bb \<in> B" using B by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1478
  def F \<equiv> "\<lambda> (f1::'a \<Rightarrow> 'b) a. if a \<in> A2 then (if a \<in> A1 then f1 a else bb)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1479
                                                else undefined"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1480
  show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1481
    show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1482
      fix f g assume f: "f \<in> Func A1 B" and g: "g \<in> Func A1 B" and eq: "F f = F g"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1483
      show "f = g"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1484
      proof(rule ext)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1485
        fix a show "f a = g a"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1486
        proof(cases "a \<in> A1")
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1487
          case True
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1488
          thus ?thesis using eq A12 unfolding F_def fun_eq_iff
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1489
          by (elim allE[of _ a]) auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1490
        qed(insert f g, unfold Func_def, fastforce)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1491
      qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1492
    qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1493
  qed(insert bb, unfold Func_def F_def, force)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1494
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1495
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1496
lemma card_of_Func_option_mono:
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1497
fixes A1 A2 :: "'a set" and B :: "'b set"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1498
assumes A12: "A1 \<subseteq> A2" and B: "B \<noteq> {}"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1499
shows "|Func_option A1 B| \<le>o |Func_option A2 B|"
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1500
by (metis card_of_Func_mono[OF A12 B] card_of_Func_option_Func
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1501
  ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1502
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1503
lemma card_of_Pfunc_Pow_Func_option:
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1504
assumes "B \<noteq> {}"
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 60585
diff changeset
  1505
shows "|Pfunc A B| \<le>o |Pow A \<times> Func_option A B|"
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1506
proof-
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1507
  have "|Pfunc A B| =o |\<Union>A' \<in> Pow A. Func_option A' B|" (is "_ =o ?K")
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1508
    unfolding Pfunc_Func_option by(rule card_of_refl)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1509
  also have "?K \<le>o |Sigma (Pow A) (\<lambda> A'. Func_option A' B)|" using card_of_UNION_Sigma .
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 60585
diff changeset
  1510
  also have "|Sigma (Pow A) (\<lambda> A'. Func_option A' B)| \<le>o |Pow A \<times> Func_option A B|"
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1511
    apply(rule card_of_Sigma_mono1) using card_of_Func_option_mono[OF _ assms] by auto
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1512
  finally show ?thesis .
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1513
qed
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1514
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1515
lemma Bpow_ordLeq_Func_Field:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1516
assumes rc: "Card_order r" and r: "Field r \<noteq> {}" and A: "\<not>finite A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1517
shows "|Bpow r A| \<le>o |Func (Field r) A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1518
proof-
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1519
  let ?F = "\<lambda> f. {x | x a. f a = x \<and> a \<in> Field r}"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1520
  {fix X assume "X \<in> Bpow r A - {{}}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1521
   hence XA: "X \<subseteq> A" and "|X| \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1522
   and X: "X \<noteq> {}" unfolding Bpow_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1523
   hence "|X| \<le>o |Field r|" by (metis Field_card_of card_of_mono2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1524
   then obtain F where 1: "X = F ` (Field r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1525
   using card_of_ordLeq2[OF X] by metis
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1526
   def f \<equiv> "\<lambda> i. if i \<in> Field r then F i else undefined"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1527
   have "\<exists> f \<in> Func (Field r) A. X = ?F f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1528
   apply (intro bexI[of _ f]) using 1 XA unfolding Func_def f_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1529
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1530
  hence "Bpow r A - {{}} \<subseteq> ?F ` (Func (Field r) A)" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1531
  hence "|Bpow r A - {{}}| \<le>o |Func (Field r) A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1532
  by (rule surj_imp_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1533
  moreover
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1534
  {have 2: "\<not>finite (Bpow r A)" using infinite_Bpow[OF rc r A] .
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1535
   have "|Bpow r A| =o |Bpow r A - {{}}|"
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1536
     by (metis 2 infinite_card_of_diff_singl ordIso_symmetric)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1537
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1538
  ultimately show ?thesis by (metis ordIso_ordLeq_trans)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1539
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1540
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1541
lemma empty_in_Func[simp]:
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
  1542
"B \<noteq> {} \<Longrightarrow> (\<lambda>x. undefined) \<in> Func {} B"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1543
unfolding Func_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1544
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1545
lemma Func_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1546
assumes "B1 \<subseteq> B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1547
shows "Func A B1 \<subseteq> Func A B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1548
using assms unfolding Func_def by force
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1549
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1550
lemma Pfunc_mono[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1551
assumes "A1 \<subseteq> A2" and "B1 \<subseteq> B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1552
shows "Pfunc A B1 \<subseteq> Pfunc A B2"
59166
4e43651235b2 recovered metis proof after 115965966e15 (Odd clash of type variables!?);
wenzelm
parents: 58889
diff changeset
  1553
using assms unfolding Pfunc_def
4e43651235b2 recovered metis proof after 115965966e15 (Odd clash of type variables!?);
wenzelm
parents: 58889
diff changeset
  1554
apply safe
4e43651235b2 recovered metis proof after 115965966e15 (Odd clash of type variables!?);
wenzelm
parents: 58889
diff changeset
  1555
apply (case_tac "x a", auto)
4e43651235b2 recovered metis proof after 115965966e15 (Odd clash of type variables!?);
wenzelm
parents: 58889
diff changeset
  1556
apply (metis in_mono option.simps(5))
4e43651235b2 recovered metis proof after 115965966e15 (Odd clash of type variables!?);
wenzelm
parents: 58889
diff changeset
  1557
done
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1558
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1559
lemma card_of_Func_UNIV_UNIV:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1560
"|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \<Rightarrow> 'b) set|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1561
using card_of_Func_UNIV[of "UNIV::'b set"] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1562
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1563
lemma ordLeq_Func:
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1564
assumes "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1565
shows "|A| \<le>o |Func A B|"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1566
unfolding card_of_ordLeq[symmetric] proof(intro exI conjI)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1567
  let ?F = "\<lambda> aa a. if a \<in> A then (if a = aa then b1 else b2) else undefined"
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1568
  show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1569
  show "?F ` A \<subseteq> Func A B" using assms unfolding Func_def by auto
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1570
qed
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1571
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1572
lemma infinite_Func:
54578
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1573
assumes A: "\<not>finite A" and B: "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
9387251b6a46 eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents: 54481
diff changeset
  1574
shows "\<not>finite (Func A B)"
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1575
using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite)
b4d644be252c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
  1576
55102
761e40ce91bc whitespace tuning
blanchet
parents: 55087
diff changeset
  1577
761e40ce91bc whitespace tuning
blanchet
parents: 55087
diff changeset
  1578
subsection {* Infinite cardinals are limit ordinals *}
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1579
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1580
lemma card_order_infinite_isLimOrd:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1581
assumes c: "Card_order r" and i: "\<not>finite (Field r)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1582
shows "isLimOrd r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1583
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1584
  have 0: "wo_rel r" and 00: "Well_order r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1585
  using c unfolding card_order_on_def wo_rel_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1586
  hence rr: "Refl r" by (metis wo_rel.REFL)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1587
  show ?thesis unfolding wo_rel.isLimOrd_def[OF 0] wo_rel.isSuccOrd_def[OF 0] proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1588
    fix j assume j: "j \<in> Field r" and jm: "\<forall>i\<in>Field r. (i, j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1589
    def p \<equiv> "Restr r (Field r - {j})"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1590
    have fp: "Field p = Field r - {j}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1591
    unfolding p_def apply(rule Refl_Field_Restr2[OF rr]) by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1592
    have of: "ofilter r (Field p)" unfolding wo_rel.ofilter_def[OF 0] proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1593
      fix a x assume a: "a \<in> Field p" and "x \<in> under r a"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1594
      hence x: "(x,a) \<in> r" "x \<in> Field r" unfolding under_def Field_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1595
      moreover have a: "a \<noteq> j" "a \<in> Field r" "(a,j) \<in> r" using a jm  unfolding fp by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1596
      ultimately have "x \<noteq> j" using j jm  by (metis 0 wo_rel.max2_def wo_rel.max2_equals1)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1597
      thus "x \<in> Field p" using x unfolding fp by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1598
    qed(unfold p_def Field_def, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1599
    have "p <o r" using j ofilter_ordLess[OF 00 of] unfolding fp p_def[symmetric] by auto
55056
b5c94200d081 renamed '_FP' files to 'BNF_' files
blanchet
parents: 55023
diff changeset
  1600
    hence 2: "|Field p| <o r" using c by (metis BNF_Cardinal_Order_Relation.ordLess_Field)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1601
    have "|Field p| =o |Field r|" unfolding fp using i by (metis infinite_card_of_diff_singl)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1602
    also have "|Field r| =o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1603
    using c by (metis card_of_unique ordIso_symmetric)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1604
    finally have "|Field p| =o r" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1605
    with 2 show False by (metis not_ordLess_ordIso)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1606
  qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1607
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1608
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1609
lemma insert_Chain:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1610
assumes "Refl r" "C \<in> Chains r" and "i \<in> Field r" and "\<And>j. j \<in> C \<Longrightarrow> (j,i) \<in> r \<or> (i,j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1611
shows "insert i C \<in> Chains r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1612
using assms unfolding Chains_def by (auto dest: refl_onD)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1613
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1614
lemma Collect_insert: "{R j |j. j \<in> insert j1 J} = insert (R j1) {R j |j. j \<in> J}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1615
by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1616
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1617
lemma Field_init_seg_of[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1618
"Field init_seg_of = UNIV"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1619
unfolding Field_def init_seg_of_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1620
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1621
lemma refl_init_seg_of[intro, simp]: "refl init_seg_of"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1622
unfolding refl_on_def Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1623
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1624
lemma regularCard_all_ex:
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1625
assumes r: "Card_order r"   "regularCard r"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1626
and As: "\<And> i j b. b \<in> B \<Longrightarrow> (i,j) \<in> r \<Longrightarrow> P i b \<Longrightarrow> P j b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1627
and Bsub: "\<forall> b \<in> B. \<exists> i \<in> Field r. P i b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1628
and cardB: "|B| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1629
shows "\<exists> i \<in> Field r. \<forall> b \<in> B. P i b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1630
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1631
  let ?As = "\<lambda>i. {b \<in> B. P i b}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1632
  have "EX i : Field r. B \<le> ?As i"
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1633
  apply(rule regularCard_UNION) using assms unfolding relChain_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1634
  thus ?thesis by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1635
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1636
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1637
lemma relChain_stabilize:
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1638
assumes rc: "relChain r As" and AsB: "(\<Union>i \<in> Field r. As i) \<subseteq> B" and Br: "|B| <o r"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1639
and ir: "\<not>finite (Field r)" and cr: "Card_order r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1640
shows "\<exists> i \<in> Field r. As (succ r i) = As i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1641
proof(rule ccontr, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1642
  have 0: "wo_rel r" and 00: "Well_order r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1643
  unfolding wo_rel_def by (metis card_order_on_well_order_on cr)+
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1644
  have L: "isLimOrd r" using ir cr by (metis card_order_infinite_isLimOrd)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1645
  have AsBs: "(\<Union>i \<in> Field r. As (succ r i)) \<subseteq> B"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1646
  using AsB L apply safe by (metis "0" UN_I set_mp wo_rel.isLimOrd_succ)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1647
  assume As_s: "\<forall>i\<in>Field r. As (succ r i) \<noteq> As i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1648
  have 1: "\<forall>i j. (i,j) \<in> r \<and> i \<noteq> j \<longrightarrow> As i \<subset> As j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1649
  proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1650
    fix i j assume 1: "(i, j) \<in> r" "i \<noteq> j" and Asij: "As i = As j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1651
    hence rij: "(succ r i, j) \<in> r" by (metis "0" wo_rel.succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1652
    hence "As (succ r i) \<subseteq> As j" using rc unfolding relChain_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1653
    moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1654
    {have "(i,succ r i) \<in> r" apply(rule wo_rel.succ_in[OF 0])
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1655
     using 1 unfolding aboveS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1656
     hence "As i \<subset> As (succ r i)" using As_s rc rij unfolding relChain_def Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1657
    }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1658
    ultimately show False unfolding Asij by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1659
  qed (insert rc, unfold relChain_def, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1660
  hence "\<forall> i \<in> Field r. \<exists> a. a \<in> As (succ r i) - As i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1661
  using wo_rel.succ_in[OF 0] AsB
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1662
  by(metis 0 card_order_infinite_isLimOrd cr ir psubset_imp_ex_mem
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1663
            wo_rel.isLimOrd_aboveS wo_rel.succ_diff)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1664
  then obtain f where f: "\<forall> i \<in> Field r. f i \<in> As (succ r i) - As i" by metis
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1665
  have "inj_on f (Field r)" unfolding inj_on_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1666
    fix i j assume ij: "i \<in> Field r" "j \<in> Field r" and fij: "f i = f j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1667
    show "i = j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1668
    proof(cases rule: wo_rel.cases_Total3[OF 0], safe)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1669
      assume "(i, j) \<in> r" and ijd: "i \<noteq> j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1670
      hence rij: "(succ r i, j) \<in> r" by (metis "0" wo_rel.succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1671
      hence "As (succ r i) \<subseteq> As j" using rc unfolding relChain_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1672
      thus "i = j" using ij ijd fij f by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1673
    next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1674
      assume "(j, i) \<in> r" and ijd: "i \<noteq> j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1675
      hence rij: "(succ r j, i) \<in> r" by (metis "0" wo_rel.succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1676
      hence "As (succ r j) \<subseteq> As i" using rc unfolding relChain_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1677
      thus "j = i" using ij ijd fij f by fastforce
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1678
    qed(insert ij, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1679
  qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1680
  moreover have "f ` (Field r) \<subseteq> B" using f AsBs by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1681
  moreover have "|B| <o |Field r|" using Br cr by (metis card_of_unique ordLess_ordIso_trans)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1682
  ultimately show False unfolding card_of_ordLess[symmetric] by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1683
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1684
55102
761e40ce91bc whitespace tuning
blanchet
parents: 55087
diff changeset
  1685
761e40ce91bc whitespace tuning
blanchet
parents: 55087
diff changeset
  1686
subsection {* Regular vs. stable cardinals *}
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1687
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1688
definition stable :: "'a rel \<Rightarrow> bool"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1689
where
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1690
"stable r \<equiv> \<forall>(A::'a set) (F :: 'a \<Rightarrow> 'a set).
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1691
               |A| <o r \<and> (\<forall>a \<in> A. |F a| <o r)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1692
               \<longrightarrow> |SIGMA a : A. F a| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1693
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1694
lemma regularCard_stable:
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1695
assumes cr: "Card_order r" and ir: "\<not>finite (Field r)" and reg: "regularCard r"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1696
shows "stable r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1697
unfolding stable_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1698
  fix A :: "'a set" and F :: "'a \<Rightarrow> 'a set" assume A: "|A| <o r" and F: "\<forall>a\<in>A. |F a| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1699
  {assume "r \<le>o |Sigma A F|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1700
   hence "|Field r| \<le>o |Sigma A F|" using card_of_Field_ordIso[OF cr]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1701
   by (metis Field_card_of card_of_cong ordLeq_iff_ordLess_or_ordIso ordLeq_ordLess_trans)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1702
   moreover have Fi: "Field r \<noteq> {}" using ir by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1703
   ultimately obtain f where f: "f ` Sigma A F = Field r" using card_of_ordLeq2 by metis
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1704
   have r: "wo_rel r" using cr unfolding card_order_on_def wo_rel_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1705
   {fix a assume a: "a \<in> A"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1706
    def L \<equiv> "{(a,u) | u. u \<in> F a}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1707
    have fL: "f ` L \<subseteq> Field r" using f a unfolding L_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1708
    have "|L| =o |F a|" unfolding L_def card_of_ordIso[symmetric]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1709
    apply(rule exI[of _ snd]) unfolding bij_betw_def inj_on_def by (auto simp: image_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1710
    hence "|L| <o r" using F a ordIso_ordLess_trans[of "|L|" "|F a|"] unfolding L_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1711
    hence "|f ` L| <o r" using ordLeq_ordLess_trans[OF card_of_image, of "L"] unfolding L_def by auto
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1712
    hence "\<not> cofinal (f ` L) r" using reg fL unfolding regularCard_def by (metis not_ordLess_ordIso)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1713
    then obtain k where k: "k \<in> Field r" and "\<forall> l \<in> L. \<not> (f l \<noteq> k \<and> (k, f l) \<in> r)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1714
    unfolding cofinal_def image_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1715
    hence "\<exists> k \<in> Field r. \<forall> l \<in> L. (f l, k) \<in> r" using r by (metis fL image_subset_iff wo_rel.in_notinI)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1716
    hence "\<exists> k \<in> Field r. \<forall> u \<in> F a. (f (a,u), k) \<in> r" unfolding L_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1717
   }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1718
   then obtain gg where gg: "\<forall> a \<in> A. \<forall> u \<in> F a. (f (a,u), gg a) \<in> r" by metis
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1719
   obtain j0 where j0: "j0 \<in> Field r" using Fi by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1720
   def g \<equiv> "\<lambda> a. if F a \<noteq> {} then gg a else j0"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1721
   have g: "\<forall> a \<in> A. \<forall> u \<in> F a. (f (a,u),g a) \<in> r" using gg unfolding g_def by auto
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1722
   hence 1: "Field r \<subseteq> (\<Union>a \<in> A. under r (g a))"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1723
   using f[symmetric] unfolding under_def image_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1724
   have gA: "g ` A \<subseteq> Field r" using gg j0 unfolding Field_def g_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1725
   moreover have "cofinal (g ` A) r" unfolding cofinal_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1726
     fix i assume "i \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1727
     then obtain j where ij: "(i,j) \<in> r" "i \<noteq> j" using cr ir by (metis infinite_Card_order_limit)
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1728
     hence "j \<in> Field r" by (metis card_order_on_def cr well_order_on_domain)
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54989
diff changeset
  1729
     then obtain a where a: "a \<in> A" and j: "(j, g a) \<in> r" using 1 unfolding under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1730
     hence "(i, g a) \<in> r" using ij wo_rel.TRANS[OF r] unfolding trans_def by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1731
     moreover have "i \<noteq> g a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1732
     using ij j r unfolding wo_rel_def unfolding well_order_on_def linear_order_on_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1733
     partial_order_on_def antisym_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1734
     ultimately show "\<exists>j\<in>g ` A. i \<noteq> j \<and> (i, j) \<in> r" using a by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1735
   qed
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1736
   ultimately have "|g ` A| =o r" using reg unfolding regularCard_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1737
   moreover have "|g ` A| \<le>o |A|" by (metis card_of_image)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1738
   ultimately have False using A by (metis not_ordLess_ordIso ordLeq_ordLess_trans)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1739
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1740
  thus "|Sigma A F| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1741
  using cr not_ordLess_iff_ordLeq by (metis card_of_Well_order card_order_on_well_order_on)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1742
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1743
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1744
lemma stable_regularCard:
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1745
assumes cr: "Card_order r" and ir: "\<not>finite (Field r)" and st: "stable r"
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1746
shows "regularCard r"
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1747
unfolding regularCard_def proof safe
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1748
  fix K assume K: "K \<subseteq> Field r" and cof: "cofinal K r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1749
  have "|K| \<le>o r" using K by (metis card_of_Field_ordIso card_of_mono1 cr ordLeq_ordIso_trans)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1750
  moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1751
  {assume Kr: "|K| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1752
   then obtain f where "\<forall> a \<in> Field r. f a \<in> K \<and> a \<noteq> f a \<and> (a, f a) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1753
   using cof unfolding cofinal_def by metis
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1754
   hence "Field r \<subseteq> (\<Union>a \<in> K. underS r a)" unfolding underS_def by auto
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1755
   hence "r \<le>o |\<Union>a \<in> K. underS r a|" using cr
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1756
   by (metis Card_order_iff_ordLeq_card_of card_of_mono1 ordLeq_transitive)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1757
   also have "|\<Union>a \<in> K. underS r a| \<le>o |SIGMA a: K. underS r a|" by (rule card_of_UNION_Sigma)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1758
   also
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1759
   {have "\<forall> a \<in> K. |underS r a| <o r" using K by (metis card_of_underS cr subsetD)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1760
    hence "|SIGMA a: K. underS r a| <o r" using st Kr unfolding stable_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1761
   }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1762
   finally have "r <o r" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1763
   hence False by (metis ordLess_irreflexive)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1764
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1765
  ultimately show "|K| =o r" by (metis ordLeq_iff_ordLess_or_ordIso)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1766
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1767
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1768
(* Note that below the types of A and F are now unconstrained: *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1769
lemma stable_elim:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1770
assumes ST: "stable r" and A_LESS: "|A| <o r" and
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1771
        F_LESS: "\<And> a. a \<in> A \<Longrightarrow> |F a| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1772
shows "|SIGMA a : A. F a| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1773
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1774
  obtain A' where 1: "A' < Field r \<and> |A'| <o r" and 2: " |A| =o |A'|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1775
  using internalize_card_of_ordLess[of A r] A_LESS by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1776
  then obtain G where 3: "bij_betw G A' A"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1777
  using card_of_ordIso  ordIso_symmetric by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1778
  (*  *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1779
  {fix a assume "a \<in> A"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1780
   hence "\<exists>B'. B' \<le> Field r \<and> |F a| =o |B'| \<and> |B'| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1781
   using internalize_card_of_ordLess[of "F a" r] F_LESS by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1782
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1783
  then obtain F' where
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1784
  temp: "\<forall>a \<in> A. F' a \<le> Field r \<and> |F a| =o |F' a| \<and> |F' a| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1785
  using bchoice[of A "\<lambda> a B'. B' \<le> Field r \<and> |F a| =o |B'| \<and> |B'| <o r"] by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1786
  hence 4: "\<forall>a \<in> A. F' a \<le> Field r \<and> |F' a| <o r" by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1787
  have 5: "\<forall>a \<in> A. |F' a| =o |F a|" using temp ordIso_symmetric by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1788
  (*  *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1789
  have "\<forall>a' \<in> A'. F'(G a') \<le> Field r \<and> |F'(G a')| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1790
  using 3 4 bij_betw_ball[of G A' A] by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1791
  hence "|SIGMA a' : A'. F'(G a')| <o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1792
  using ST 1 unfolding stable_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1793
  moreover have "|SIGMA a' : A'. F'(G a')| =o |SIGMA a : A. F a|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1794
  using card_of_Sigma_cong[of G A' A F' F] 5 3 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1795
  ultimately show ?thesis using ordIso_symmetric ordIso_ordLess_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1796
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1797
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1798
lemma stable_natLeq: "stable natLeq"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1799
proof(unfold stable_def, safe)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1800
  fix A :: "'a set" and F :: "'a \<Rightarrow> 'a set"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1801
  assume "|A| <o natLeq" and "\<forall>a\<in>A. |F a| <o natLeq"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1802
  hence "finite A \<and> (\<forall>a \<in> A. finite(F a))"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1803
  by (auto simp add: finite_iff_ordLess_natLeq)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1804
  hence "finite(Sigma A F)" by (simp only: finite_SigmaI[of A F])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1805
  thus "|Sigma A F | <o natLeq"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1806
  by (auto simp add: finite_iff_ordLess_natLeq)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1807
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1808
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1809
corollary regularCard_natLeq: "regularCard natLeq"
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1810
using stable_regularCard[OF natLeq_Card_order _ stable_natLeq] Field_natLeq by simp
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1811
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1812
lemma stable_cardSuc:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1813
assumes CARD: "Card_order r" and INF: "\<not>finite (Field r)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1814
shows "stable(cardSuc r)"
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1815
using infinite_cardSuc_regularCard regularCard_stable
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1816
by (metis CARD INF cardSuc_Card_order cardSuc_finite)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1817
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1818
lemma stable_UNION:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1819
assumes ST: "stable r" and A_LESS: "|A| <o r" and
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1820
        F_LESS: "\<And> a. a \<in> A \<Longrightarrow> |F a| <o r"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1821
shows "|\<Union>a \<in> A. F a| <o r"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1822
proof-
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 59166
diff changeset
  1823
  have "|\<Union>a \<in> A. F a| \<le>o |SIGMA a : A. F a|"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1824
  using card_of_UNION_Sigma by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1825
  thus ?thesis using assms stable_elim ordLeq_ordLess_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1826
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1827
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1828
lemma stable_ordIso1:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1829
assumes ST: "stable r" and ISO: "r' =o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1830
shows "stable r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1831
proof(unfold stable_def, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1832
  fix A::"'b set" and F::"'b \<Rightarrow> 'b set"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1833
  assume "|A| <o r'" and "\<forall>a \<in> A. |F a| <o r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1834
  hence "( |A| <o r) \<and> (\<forall>a \<in> A. |F a| <o r)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1835
  using ISO ordLess_ordIso_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1836
  hence "|SIGMA a : A. F a| <o r" using assms stable_elim by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1837
  thus "|SIGMA a : A. F a| <o r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1838
  using ISO ordIso_symmetric ordLess_ordIso_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1839
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1840
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1841
lemma stable_ordIso2:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1842
assumes ST: "stable r" and ISO: "r =o r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1843
shows "stable r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1844
using assms stable_ordIso1 ordIso_symmetric by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1845
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1846
lemma stable_ordIso:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1847
assumes "r =o r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1848
shows "stable r = stable r'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1849
using assms stable_ordIso1 stable_ordIso2 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1850
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1851
lemma stable_nat: "stable |UNIV::nat set|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1852
using stable_natLeq card_of_nat stable_ordIso by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1853
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1854
text{* Below, the type of "A" is not important -- we just had to choose an appropriate
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1855
   type to make "A" possible. What is important is that arbitrarily large
54989
0fd6b0660242 shot in the dark to make LaTeX happy again
haftmann
parents: 54980
diff changeset
  1856
   infinite sets of stable cardinality exist. *}
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1857
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1858
lemma infinite_stable_exists:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1859
assumes CARD: "\<forall>r \<in> R. Card_order (r::'a rel)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1860
shows "\<exists>(A :: (nat + 'a set)set).
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1861
          \<not>finite A \<and> stable |A| \<and> (\<forall>r \<in> R. r <o |A| )"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1862
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1863
  have "\<exists>(A :: (nat + 'a set)set).
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1864
          \<not>finite A \<and> stable |A| \<and> |UNIV::'a set| <o |A|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1865
  proof(cases "finite (UNIV::'a set)")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1866
    assume Case1: "finite (UNIV::'a set)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1867
    let ?B = "UNIV::nat set"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1868
    have "\<not>finite(?B <+> {})" using finite_Plus_iff by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1869
    moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1870
    have "stable |?B|" using stable_natLeq card_of_nat stable_ordIso1 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1871
    hence "stable |?B <+> {}|" using stable_ordIso card_of_Plus_empty1 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1872
    moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1873
    have "\<not>finite(Field |?B| ) \<and> finite(Field |UNIV::'a set| )" using Case1 by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1874
    hence "|UNIV::'a set| <o |?B|" by (simp add: finite_ordLess_infinite)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1875
    hence "|UNIV::'a set| <o |?B <+> {}|" using card_of_Plus_empty1 ordLess_ordIso_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1876
    ultimately show ?thesis by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1877
  next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1878
    assume Case1: "\<not>finite (UNIV::'a set)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1879
    hence *: "\<not>finite(Field |UNIV::'a set| )" by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1880
    let ?B = "Field(cardSuc |UNIV::'a set| )"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1881
    have 0: "|?B| =o |{} <+> ?B|" using card_of_Plus_empty2 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1882
    have 1: "\<not>finite ?B" using Case1 card_of_cardSuc_finite by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1883
    hence 2: "\<not>finite({} <+> ?B)" using 0 card_of_ordIso_finite by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1884
    have "|?B| =o cardSuc |UNIV::'a set|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1885
    using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1886
    moreover have "stable(cardSuc |UNIV::'a set| )"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1887
    using stable_cardSuc * card_of_Card_order by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1888
    ultimately have "stable |?B|" using stable_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1889
    hence 3: "stable |{} <+> ?B|" using stable_ordIso 0 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1890
    have "|UNIV::'a set| <o cardSuc |UNIV::'a set|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1891
    using card_of_Card_order cardSuc_greater by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1892
    moreover have "|?B| =o  cardSuc |UNIV::'a set|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1893
    using card_of_Card_order cardSuc_Card_order  card_of_Field_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1894
    ultimately have "|UNIV::'a set| <o |?B|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1895
    using ordIso_symmetric ordLess_ordIso_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1896
    hence "|UNIV::'a set| <o |{} <+> ?B|" using 0 ordLess_ordIso_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1897
    thus ?thesis using 2 3 by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1898
  qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1899
  thus ?thesis using CARD card_of_UNIV2 ordLeq_ordLess_trans by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1900
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1901
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1902
corollary infinite_regularCard_exists:
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1903
assumes CARD: "\<forall>r \<in> R. Card_order (r::'a rel)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1904
shows "\<exists>(A :: (nat + 'a set)set).
55087
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1905
          \<not>finite A \<and> regularCard |A| \<and> (\<forall>r \<in> R. r <o |A| )"
252c7fec4119 renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents: 55056
diff changeset
  1906
using infinite_stable_exists[OF CARD] stable_regularCard by (metis Field_card_of card_of_card_order_on)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
  1907
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1908
end