src/HOL/Cardinals/Cardinal_Arithmetic.thy
author wenzelm
Mon, 24 Oct 2022 20:37:32 +0200
changeset 76371 1ac2416e8432
parent 75132 e349c2da30d2
child 76946 5df58a471d9e
permissions -rw-r--r--
tuned signature (again, amending f32ac01aef5e), e.g. relevant for Isabelle/DOF;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 48975
diff changeset
     1
(*  Title:      HOL/Cardinals/Cardinal_Arithmetic.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    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 arithmetic.
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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
     8
section \<open>Cardinal Arithmetic\<close>
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_Arithmetic
68652
1e37b45ce3fb normalize imports
traytel
parents: 66453
diff changeset
    11
imports Cardinal_Order_Relation
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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
    14
subsection \<open>Binary sum\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
lemma csum_Cnotzero2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  "Cnotzero r2 \<Longrightarrow> Cnotzero (r1 +c r2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
unfolding csum_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
lemma single_cone:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
  "|{x}| =o cone"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
proof -
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
  let ?f = "\<lambda>x. ()"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
  have "bij_betw ?f {x} {()}" unfolding bij_betw_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
  thus ?thesis unfolding cone_def 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
    27
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
lemma cone_Cnotzero: "Cnotzero cone"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
by (simp add: cone_not_czero Card_order_cone)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
lemma cone_ordLeq_ctwo: "cone \<le>o ctwo"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    35
lemma csum_czero1: "Card_order r \<Longrightarrow> r +c czero =o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    36
  unfolding czero_def csum_def Field_card_of
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    37
  by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty1] card_of_Field_ordIso])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    38
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    39
lemma csum_czero2: "Card_order r \<Longrightarrow> czero +c r =o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    40
  unfolding czero_def csum_def Field_card_of
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    41
  by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty2] card_of_Field_ordIso])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    42
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
    44
subsection \<open>Product\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
lemma Times_cprod: "|A \<times> B| =o |A| *c |B|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
by (simp only: cprod_def Field_card_of card_of_refl)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    49
lemma card_of_Times_singleton:
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
    50
  fixes A :: "'a set"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
    51
  shows "|A \<times> {x}| =o |A|"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    52
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
    53
  define f :: "'a \<times> 'b \<Rightarrow> 'a" where "f = (\<lambda>(a, b). a)"
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 60585
diff changeset
    54
  have "A \<subseteq> f ` (A \<times> {x})" unfolding f_def by (auto simp: image_iff)
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 60585
diff changeset
    55
  hence "bij_betw f (A \<times> {x}) A"  unfolding bij_betw_def inj_on_def f_def by fastforce
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    56
  thus ?thesis using card_of_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    57
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    58
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    59
lemma cprod_assoc: "(r *c s) *c t =o r *c s *c t"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    60
  unfolding cprod_def Field_card_of by (rule card_of_Times_assoc)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    61
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    62
lemma cprod_czero: "r *c czero =o czero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    63
  unfolding cprod_def czero_def Field_card_of by (simp add: card_of_empty_ordIso)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    64
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    65
lemma cprod_cone: "Card_order r \<Longrightarrow> r *c cone =o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    66
  unfolding cprod_def cone_def Field_card_of
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    67
  by (drule card_of_Field_ordIso) (erule ordIso_transitive[OF card_of_Times_singleton])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
    68
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
lemma ordLeq_cprod1: "\<lbrakk>Card_order p1; Cnotzero p2\<rbrakk> \<Longrightarrow> p1 \<le>o p1 *c p2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
unfolding cprod_def by (metis Card_order_Times1 czeroI)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
    74
subsection \<open>Exponentiation\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
lemma cexp_czero: "r ^c czero =o cone"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
lemma Pow_cexp_ctwo:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
  "|Pow A| =o ctwo ^c |A|"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
lemma Cnotzero_cexp:
75132
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
    84
  assumes "Cnotzero q" 
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
  shows "Cnotzero (q ^c r)"
75132
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
    86
proof -
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
    87
  have "Field q \<noteq> {}"
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
    88
    by (metis Card_order_iff_ordIso_card_of assms(1) czero_def)
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
    89
  then show ?thesis
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
    90
    by (simp add: card_of_ordIso_czero_iff_empty cexp_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
lemma Cinfinite_ctwo_cexp:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
  "Cinfinite r \<Longrightarrow> Cinfinite (ctwo ^c r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
unfolding ctwo_def cexp_def cinfinite_def Field_card_of
54475
b4d644be252c moved theorems out of LFP
blanchet
parents: 54474
diff changeset
    96
by (rule conjI, rule infinite_Func, auto)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
lemma cone_ordLeq_iff_Field:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
  assumes "cone \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
  shows "Field r \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
proof (rule ccontr)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
  assume "\<not> Field r \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
  hence "Field r = {}" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
  thus False using card_of_empty3
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
    card_of_mono2[OF assms] Cnotzero_imp_not_empty[OF cone_Cnotzero] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
lemma cone_ordLeq_cexp: "cone \<le>o r1 \<Longrightarrow> cone \<le>o r1 ^c r2"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   109
by (simp add: cexp_def cone_def Func_non_emp cone_ordLeq_iff_Field)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
lemma Card_order_czero: "Card_order czero"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
by (simp only: card_of_Card_order czero_def)
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 cexp_mono2'':
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
  assumes 2: "p2 \<le>o r2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
  and n1: "Cnotzero q"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
  and n2: "Card_order p2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
  shows "q ^c p2 \<le>o q ^c r2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
proof (cases "p2 =o (czero :: 'a rel)")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
  case True
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
  hence "q ^c p2 =o q ^c (czero :: 'a rel)" using n1 n2 cexp_cong2 Card_order_czero by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
  also have "q ^c (czero :: 'a rel) =o cone" using cexp_czero by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
  also have "cone \<le>o q ^c r2" using cone_ordLeq_cexp cone_ordLeq_Cnotzero n1 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
  finally show ?thesis .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
  case False thus ?thesis using assms cexp_mono2' czeroI by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
lemma csum_cexp: "\<lbrakk>Cinfinite r1; Cinfinite r2; Card_order q; ctwo \<le>o q\<rbrakk> \<Longrightarrow>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
  q ^c r1 +c q ^c r2 \<le>o q ^c (r1 +c r2)"
75132
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
   131
  apply (rule csum_cinfinite_bound)
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
   132
  apply (metis cexp_mono2' cinfinite_def finite.emptyI ordLeq_csum1)
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
   133
  apply (metis cexp_mono2' cinfinite_def finite.emptyI ordLeq_csum2)
e349c2da30d2 Simplified a couple of extremely long and ugly apply-proofs
paulson <lp15@cam.ac.uk>
parents: 68652
diff changeset
   134
  by (simp_all add: Card_order_cexp Cinfinite_csum1 Cinfinite_cexp cinfinite_cexp)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
lemma csum_cexp': "\<lbrakk>Cinfinite r; Card_order q; ctwo \<le>o q\<rbrakk> \<Longrightarrow> q +c r \<le>o q ^c r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
apply (rule csum_cinfinite_bound)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
    apply (metis Cinfinite_Cnotzero ordLeq_cexp1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
   apply (metis ordLeq_cexp2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
  apply blast+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
by (metis Cinfinite_cexp)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   142
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
lemma card_of_Sigma_ordLeq_Cinfinite:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
  "\<lbrakk>Cinfinite r; |I| \<le>o r; \<forall>i \<in> I. |A i| \<le>o r\<rbrakk> \<Longrightarrow> |SIGMA i : I. A i| \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
unfolding cinfinite_def by (blast intro: card_of_Sigma_ordLeq_infinite_Field)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
54794
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   147
lemma card_order_cexp:
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   148
  assumes "card_order r1" "card_order r2"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   149
  shows "card_order (r1 ^c r2)"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   150
proof -
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   151
  have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   152
  thus ?thesis unfolding cexp_def Func_def by simp
54794
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   153
qed
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   154
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   155
lemma Cinfinite_ordLess_cexp:
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   156
  assumes r: "Cinfinite r"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   157
  shows "r <o r ^c r"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   158
proof -
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   159
  have "r <o ctwo ^c r" using r by (simp only: ordLess_ctwo_cexp)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   160
  also have "ctwo ^c r \<le>o r ^c r"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   161
    by (rule cexp_mono1[OF ctwo_ordLeq_Cinfinite]) (auto simp: r ctwo_not_czero Card_order_ctwo)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   162
  finally show ?thesis .
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   163
qed
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   164
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   165
lemma infinite_ordLeq_cexp:
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   166
  assumes "Cinfinite r"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   167
  shows "r \<le>o r ^c r"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   168
by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]])
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   169
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   170
lemma czero_cexp: "Cnotzero r \<Longrightarrow> czero ^c r =o czero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   171
  by (drule Cnotzero_imp_not_empty) (simp add: cexp_def czero_def card_of_empty_ordIso)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   172
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55851
diff changeset
   173
lemma Func_singleton:
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   174
fixes x :: 'b and A :: "'a set"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   175
shows "|Func A {x}| =o |{x}|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   176
proof (rule ordIso_symmetric)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   177
  define f where [abs_def]: "f y a = (if y = x \<and> a \<in> A then x else undefined)" for y a
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   178
  have "Func A {x} \<subseteq> f ` {x}" unfolding f_def Func_def by (force simp: fun_eq_iff)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   179
  hence "bij_betw f {x} (Func A {x})" unfolding bij_betw_def inj_on_def f_def Func_def
62390
842917225d56 more canonical names
nipkow
parents: 61943
diff changeset
   180
    by (auto split: if_split_asm)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   181
  thus "|{x}| =o |Func A {x}|" using card_of_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   182
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   183
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   184
lemma cone_cexp: "cone ^c r =o cone"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   185
  unfolding cexp_def cone_def Field_card_of by (rule Func_singleton)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   186
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   187
lemma card_of_Func_squared:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   188
  fixes A :: "'a set"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   189
  shows "|Func (UNIV :: bool set) A| =o |A \<times> A|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   190
proof (rule ordIso_symmetric)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   191
  define f where "f = (\<lambda>(x::'a,y) b. if A = {} then undefined else if b then x else y)"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   192
  have "Func (UNIV :: bool set) A \<subseteq> f ` (A \<times> A)" unfolding f_def Func_def
62390
842917225d56 more canonical names
nipkow
parents: 61943
diff changeset
   193
    by (auto simp: image_iff fun_eq_iff split: option.splits if_split_asm) blast
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   194
  hence "bij_betw f (A \<times> A) (Func (UNIV :: bool set) A)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   195
    unfolding bij_betw_def inj_on_def f_def Func_def by (auto simp: fun_eq_iff)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   196
  thus "|A \<times> A| =o |Func (UNIV :: bool set) A|" using card_of_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   197
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   198
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   199
lemma cexp_ctwo: "r ^c ctwo =o r *c r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   200
  unfolding cexp_def ctwo_def cprod_def Field_card_of by (rule card_of_Func_squared)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   201
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   202
lemma card_of_Func_Plus:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   203
  fixes A :: "'a set" and B :: "'b set" and C :: "'c set"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   204
  shows "|Func (A <+> B) C| =o |Func A C \<times> Func B C|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   205
proof (rule ordIso_symmetric)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   206
  define f where "f = (\<lambda>(g :: 'a => 'c, h::'b \<Rightarrow> 'c) ab. case ab of Inl a \<Rightarrow> g a | Inr b \<Rightarrow> h b)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   207
  define f' where "f' = (\<lambda>(f :: ('a + 'b) \<Rightarrow> 'c). (\<lambda>a. f (Inl a), \<lambda>b. f (Inr b)))"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   208
  have "f ` (Func A C \<times> Func B C) \<subseteq> Func (A <+> B) C"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   209
    unfolding Func_def f_def by (force split: sum.splits)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   210
  moreover have "f' ` Func (A <+> B) C \<subseteq> Func A C \<times> Func B C" unfolding Func_def f'_def by force
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   211
  moreover have "\<forall>a \<in> Func A C \<times> Func B C. f' (f a) = a" unfolding f'_def f_def Func_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   212
  moreover have "\<forall>a' \<in> Func (A <+> B) C. f (f' a') = a'" unfolding f'_def f_def Func_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   213
    by (auto split: sum.splits)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   214
  ultimately have "bij_betw f (Func A C \<times> Func B C) (Func (A <+> B) C)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   215
    by (intro bij_betw_byWitness[of _ f' f])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   216
  thus "|Func A C \<times> Func B C| =o |Func (A <+> B) C|" using card_of_ordIso by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   217
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   218
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   219
lemma cexp_csum: "r ^c (s +c t) =o r ^c s *c r ^c t"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   220
  unfolding cexp_def cprod_def csum_def Field_card_of by (rule card_of_Func_Plus)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   221
54794
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   222
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   223
subsection \<open>Powerset\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
54794
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   225
definition cpow where "cpow r = |Pow (Field r)|"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   226
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   227
lemma card_order_cpow: "card_order r \<Longrightarrow> card_order (cpow r)"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   228
by (simp only: cpow_def Field_card_order Pow_UNIV card_of_card_order_on)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   229
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   230
lemma cpow_greater_eq: "Card_order r \<Longrightarrow> r \<le>o cpow r"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   231
by (rule ordLess_imp_ordLeq) (simp only: cpow_def Card_order_Pow)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   232
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   233
lemma Cinfinite_cpow: "Cinfinite r \<Longrightarrow> Cinfinite (cpow r)"
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   234
unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow)
e279c2ceb54c reduced cardinals dependencies of (co)datatypes
traytel
parents: 54581
diff changeset
   235
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
lemma Card_order_cpow: "Card_order (cpow r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
unfolding cpow_def by (rule card_of_Card_order)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   238
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
lemma cardSuc_ordLeq_cpow: "Card_order r \<Longrightarrow> cardSuc r \<le>o cpow r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
unfolding cpow_def by (metis Card_order_Pow cardSuc_ordLess_ordLeq card_of_Card_order)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   245
subsection \<open>Inverse image\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   246
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   247
lemma vimage_ordLeq:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   248
assumes "|A| \<le>o k" and "\<forall> a \<in> A. |vimage f {a}| \<le>o k" and "Cinfinite k"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   249
shows "|vimage f A| \<le>o k"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   250
proof-
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58889
diff changeset
   251
  have "vimage f A = (\<Union>a \<in> A. vimage f {a})" by auto
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58889
diff changeset
   252
  also have "|\<Union>a \<in> A. vimage f {a}| \<le>o k"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   253
  using UNION_Cinfinite_bound[OF assms] .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   254
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   255
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   256
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   257
subsection \<open>Maximum\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   258
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   259
definition cmax where
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   260
  "cmax r s =
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   261
    (if cinfinite r \<or> cinfinite s then czero +c r +c s
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   262
     else natLeq_on (max (card (Field r)) (card (Field s))) +c czero)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   263
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   264
lemma cmax_com: "cmax r s =o cmax s r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   265
  unfolding cmax_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   266
  by (auto simp: max.commute intro: csum_cong2[OF csum_com] csum_cong2[OF czero_ordIso])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   267
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   268
lemma cmax1:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   269
  assumes "Card_order r" "Card_order s" "s \<le>o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   270
  shows "cmax r s =o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   271
unfolding cmax_def proof (split if_splits, intro conjI impI)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   272
  assume "cinfinite r \<or> cinfinite s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   273
  hence Cinf: "Cinfinite r" using assms(1,3) by (metis cinfinite_mono)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   274
  have "czero +c r +c s =o r +c s" by (rule csum_czero2[OF Card_order_csum])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   275
  also have "r +c s =o r" by (rule csum_absorb1[OF Cinf assms(3)])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   276
  finally show "czero +c r +c s =o r" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   277
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   278
  assume "\<not> (cinfinite r \<or> cinfinite s)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   279
  hence fin: "finite (Field r)" and "finite (Field s)" unfolding cinfinite_def by simp_all
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   280
  moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   281
  { from assms(2) have "|Field s| =o s" by (rule card_of_Field_ordIso)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   282
    also from assms(3) have "s \<le>o r" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   283
    also from assms(1) have "r =o |Field r|" by (rule ordIso_symmetric[OF card_of_Field_ordIso])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   284
    finally have "|Field s| \<le>o |Field r|" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   285
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   286
  ultimately have "card (Field s) \<le> card (Field r)" by (subst sym[OF finite_card_of_iff_card2])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   287
  hence "max (card (Field r)) (card (Field s)) = card (Field r)" by (rule max_absorb1)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   288
  hence "natLeq_on (max (card (Field r)) (card (Field s))) +c czero =
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   289
    natLeq_on (card (Field r)) +c czero" by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   290
  also have "\<dots> =o natLeq_on (card (Field r))" by (rule csum_czero1[OF natLeq_on_Card_order])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   291
  also have "natLeq_on (card (Field r)) =o |Field r|"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   292
    by (rule ordIso_symmetric[OF finite_imp_card_of_natLeq_on[OF fin]])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   293
  also from assms(1) have "|Field r| =o r" by (rule card_of_Field_ordIso)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   294
  finally show "natLeq_on (max (card (Field r)) (card (Field s))) +c czero =o r" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   295
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   296
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   297
lemma cmax2:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   298
  assumes "Card_order r" "Card_order s" "r \<le>o s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   299
  shows "cmax r s =o s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   300
  by (metis assms cmax1 cmax_com ordIso_transitive)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   301
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   302
lemma csum_absorb2: "Cinfinite r2 \<Longrightarrow> r1 \<le>o r2 \<Longrightarrow> r1 +c r2 =o r2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   303
  by (metis csum_absorb2')
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   304
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   305
lemma cprod_infinite2': "\<lbrakk>Cnotzero r1; Cinfinite r2; r1 \<le>o r2\<rbrakk> \<Longrightarrow> r1 *c r2 =o r2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   306
  unfolding ordIso_iff_ordLeq
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   307
  by (intro conjI cprod_cinfinite_bound ordLeq_cprod2 ordLeq_refl)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   308
    (auto dest!: ordIso_imp_ordLeq not_ordLeq_ordLess simp: czero_def Card_order_empty)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   309
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   310
context
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   311
  fixes r s
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   312
  assumes r: "Cinfinite r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   313
  and     s: "Cinfinite s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   314
begin
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   315
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   316
lemma cmax_csum: "cmax r s =o r +c s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   317
proof (cases "r \<le>o s")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   318
  case True
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   319
  hence "cmax r s =o s" by (metis cmax2 r s)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   320
  also have "s =o r +c s" by (metis True csum_absorb2 ordIso_symmetric s)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   321
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   322
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   323
  case False
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   324
  hence "s \<le>o r" by (metis ordLeq_total r s card_order_on_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   325
  hence "cmax r s =o r" by (metis cmax1 r s)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   326
  also have "r =o r +c s" by (metis \<open>s \<le>o r\<close> csum_absorb1 ordIso_symmetric r)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   327
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   328
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   329
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   330
lemma cmax_cprod: "cmax r s =o r *c s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   331
proof (cases "r \<le>o s")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   332
  case True
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   333
  hence "cmax r s =o s" by (metis cmax2 r s)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   334
  also have "s =o r *c s" by (metis Cinfinite_Cnotzero True cprod_infinite2' ordIso_symmetric r s)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   335
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   336
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   337
  case False
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   338
  hence "s \<le>o r" by (metis ordLeq_total r s card_order_on_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   339
  hence "cmax r s =o r" by (metis cmax1 r s)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   340
  also have "r =o r *c s" by (metis Cinfinite_Cnotzero \<open>s \<le>o r\<close> cprod_infinite1' ordIso_symmetric r s)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   341
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   342
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   343
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
end
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   345
55174
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   346
lemma Card_order_cmax:
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   347
assumes r: "Card_order r" and s: "Card_order s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   348
shows "Card_order (cmax r s)"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   349
unfolding cmax_def by (auto simp: Card_order_csum)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   350
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   351
lemma ordLeq_cmax:
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   352
assumes r: "Card_order r" and s: "Card_order s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   353
shows "r \<le>o cmax r s \<and> s \<le>o cmax r s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   354
proof-
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   355
  {assume "r \<le>o s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   356
   hence ?thesis by (metis cmax2 ordIso_iff_ordLeq ordLeq_transitive r s)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   357
  }
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   358
  moreover
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   359
  {assume "s \<le>o r"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   360
   hence ?thesis using cmax_com by (metis cmax2 ordIso_iff_ordLeq ordLeq_transitive r s)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   361
  }
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   362
  ultimately show ?thesis using r s ordLeq_total unfolding card_order_on_def by auto
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   363
qed
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   364
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   365
lemmas ordLeq_cmax1 = ordLeq_cmax[THEN conjunct1] and
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   366
       ordLeq_cmax2 = ordLeq_cmax[THEN conjunct2]
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   367
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   368
lemma finite_cmax:
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   369
assumes r: "Card_order r" and s: "Card_order s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   370
shows "finite (Field (cmax r s)) \<longleftrightarrow> finite (Field r) \<and> finite (Field s)"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   371
proof-
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   372
  {assume "r \<le>o s"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   373
   hence ?thesis by (metis cmax2 ordIso_finite_Field ordLeq_finite_Field r s)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   374
  }
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   375
  moreover
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   376
  {assume "s \<le>o r"
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   377
   hence ?thesis by (metis cmax1 ordIso_finite_Field ordLeq_finite_Field r s)
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   378
  }
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   379
  ultimately show ?thesis using r s ordLeq_total unfolding card_order_on_def by auto
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   380
qed
2e8fe898fa71 extended cardinals library
traytel
parents: 55102
diff changeset
   381
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54794
diff changeset
   382
end