src/HOL/Cardinals/Wellorder_Constructions.thy
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 81808 f575ad8dcbe5
child 82248 e8c96013ea8a
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
     1
(*  Title:      HOL/Cardinals/Wellorder_Constructions.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
Constructions on wellorders.
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>Constructions on Wellorders\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
    10
theory Wellorder_Constructions
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    11
  imports
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    12
    Wellorder_Embedding Order_Union
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
70078
3a1b2d8c89aa bundle for cardinal syntax
traytel
parents: 69712
diff changeset
    15
unbundle cardinal_syntax
3a1b2d8c89aa bundle for cardinal syntax
traytel
parents: 69712
diff changeset
    16
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
declare
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
  ordLeq_Well_order_simp[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
  not_ordLeq_iff_ordLess[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
  not_ordLess_iff_ordLeq[simp]
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
    21
  Func_empty[simp]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
    22
  Func_is_emp[simp]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
    26
subsection \<open>Order filters versus restrictions and embeddings\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
lemma ofilter_subset_iso:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    29
  assumes WELL: "Well_order r" and
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    30
    OFA: "ofilter r A" and OFB: "ofilter r B"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    31
  shows "(A = B) = iso (Restr r A) (Restr r B) id"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    32
  using assms by (auto simp add: ofilter_subset_embedS_iso)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
    35
subsection \<open>Ordering the well-orders by existence of embeddings\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
corollary ordLeq_refl_on: "refl_on {r. Well_order r} ordLeq"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    38
  using ordLeq_reflexive unfolding ordLeq_def refl_on_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    39
  by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
corollary ordLeq_trans: "trans ordLeq"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    42
  using trans_def[of ordLeq] ordLeq_transitive by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
corollary ordLeq_preorder_on: "preorder_on {r. Well_order r} ordLeq"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    45
  by(auto simp add: preorder_on_def ordLeq_refl_on ordLeq_trans)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 76950
diff changeset
    47
corollary ordIso_subset: "ordIso \<subseteq> {r. Well_order r} \<times> {r. Well_order r}"
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 76950
diff changeset
    48
  using ordIso_reflexive unfolding refl_on_def ordIso_def by blast
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 76950
diff changeset
    49
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
corollary ordIso_refl_on: "refl_on {r. Well_order r} ordIso"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    51
  using ordIso_reflexive unfolding refl_on_def ordIso_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    52
  by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
corollary ordIso_trans: "trans ordIso"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    55
  using trans_def[of ordIso] ordIso_transitive by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
corollary ordIso_sym: "sym ordIso"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    58
  by (auto simp add: sym_def ordIso_symmetric)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
corollary ordIso_equiv: "equiv {r. Well_order r} ordIso"
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 76950
diff changeset
    61
  using ordIso_subset ordIso_refl_on ordIso_sym ordIso_trans by (intro equivI)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
54476
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
    63
lemma ordLess_Well_order_simp[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    64
  assumes "r <o r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    65
  shows "Well_order r \<and> Well_order r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    66
  using assms unfolding ordLess_def by simp
54476
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
    67
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
    68
lemma ordIso_Well_order_simp[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    69
  assumes "r =o r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    70
  shows "Well_order r \<and> Well_order r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    71
  using assms unfolding ordIso_def by simp
54476
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
    72
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
lemma ordLess_irrefl: "irrefl ordLess"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    74
  by(unfold irrefl_def, auto simp add: ordLess_irreflexive)
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 ordLess_or_ordIso:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    77
  assumes WELL: "Well_order r" and WELL': "Well_order r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    78
  shows "r <o r' \<or> r' <o r \<or> r =o r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    79
  unfolding ordLess_def ordIso_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    80
  using assms embedS_or_iso[of r r'] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
corollary ordLeq_ordLess_Un_ordIso:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    83
  "ordLeq = ordLess \<union> ordIso"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    84
  by (auto simp add: ordLeq_iff_ordLess_or_ordIso)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
lemma ordIso_or_ordLess:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    87
  assumes WELL: "Well_order r" and WELL': "Well_order r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    88
  shows "r =o r' \<or> r <o r' \<or> r' <o r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    89
  using assms ordLess_or_ordLeq ordLeq_iff_ordLess_or_ordIso by blast
48975
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
lemmas ord_trans = ordIso_transitive ordLeq_transitive ordLess_transitive
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    92
  ordIso_ordLeq_trans ordLeq_ordIso_trans
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    93
  ordIso_ordLess_trans ordLess_ordIso_trans
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    94
  ordLess_ordLeq_trans ordLeq_ordLess_trans
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
lemma ofilter_ordLeq:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    97
  assumes "Well_order r" and "ofilter r A"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
    98
  shows "Restr r A \<le>o r"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
    99
by (metis assms inf.orderE ofilter_embed ofilter_subset_ordLeq refl_on_def wo_rel.Field_ofilter wo_rel.REFL wo_rel.intro)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
corollary under_Restr_ordLeq:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   102
  "Well_order r \<Longrightarrow> Restr r (under r a) \<le>o r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   103
  by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   106
subsection \<open>Copy via direct images\<close>
48975
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 Id_dir_image: "dir_image Id f \<le> Id"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   109
  unfolding dir_image_def by auto
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 Un_dir_image:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   112
  "dir_image (r1 \<union> r2) f = (dir_image r1 f) \<union> (dir_image r2 f)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   113
  unfolding dir_image_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
lemma Int_dir_image:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   116
  assumes "inj_on f (Field r1 \<union> Field r2)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   117
  shows "dir_image (r1 Int r2) f = (dir_image r1 f) Int (dir_image r2 f)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
  show "dir_image (r1 Int r2) f \<le> (dir_image r1 f) Int (dir_image r2 f)"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   120
    using assms unfolding dir_image_def inj_on_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
  show "(dir_image r1 f) Int (dir_image r2 f) \<le> dir_image (r1 Int r2) f"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   123
    by (clarsimp simp: dir_image_def) (metis FieldI1 FieldI2 UnCI assms inj_on_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
52203
055c392e79cf fixed broken Cardinals and BNF due to Order_Union
popescua
parents: 52183
diff changeset
   126
(* More facts on ordinal sum: *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
lemma Osum_embed:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   129
  assumes FLD: "Field r Int Field r' = {}" and
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   130
    WELL: "Well_order r" and WELL': "Well_order r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   131
  shows "embed r (r Osum r') id"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
  have 1: "Well_order (r Osum r')"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   134
    using assms by (auto simp add: Osum_Well_order)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
  have "compat r (r Osum r') id"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   137
    unfolding compat_def Osum_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
  have "inj_on id (Field r)" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
  have "ofilter (r Osum r') (Field r)"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   142
    using 1 FLD
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   143
    by (auto simp add: wo_rel_def wo_rel.ofilter_def Osum_def under_def Field_iff disjoint_iff)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
  ultimately show ?thesis
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   145
    using assms by (auto simp add: embed_iff_compat_inj_on_ofilter)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   147
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   148
corollary Osum_ordLeq:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   149
  assumes FLD: "Field r Int Field r' = {}" and
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   150
    WELL: "Well_order r" and WELL': "Well_order r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   151
  shows "r \<le>o r Osum r'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   152
  using assms Osum_embed Osum_Well_order
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   153
  unfolding ordLeq_def by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
lemma Well_order_embed_copy:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   156
  assumes WELL: "well_order_on A r" and
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   157
    INJ: "inj_on f A" and SUB: "f ` A \<le> B"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   158
  shows "\<exists>r'. well_order_on B r' \<and> r \<le>o r'"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
  have "bij_betw f A (f ` A)"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   161
    using INJ inj_on_imp_bij_betw by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
  then obtain r'' where "well_order_on (f ` A) r''" and 1: "r =o r''"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   163
    using WELL  Well_order_iso_copy by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
  hence 2: "Well_order r'' \<and> Field r'' = (f ` A)"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   165
    using well_order_on_Well_order by blast
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   166
      (*  *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
  let ?C = "B - (f ` A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
  obtain r''' where "well_order_on ?C r'''"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   169
    using well_order_on by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
  hence 3: "Well_order r''' \<and> Field r''' = ?C"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   171
    using well_order_on_Well_order by blast
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   172
      (*  *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
  let ?r' = "r'' Osum r'''"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
  have "Field r'' Int Field r''' = {}"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   175
    using 2 3 by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
  hence "r'' \<le>o ?r'" using Osum_ordLeq[of r'' r'''] 2 3 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
  hence 4: "r \<le>o ?r'" using 1 ordIso_ordLeq_trans by blast
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   178
      (*  *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
  hence "Well_order ?r'" unfolding ordLeq_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
  have "Field ?r' = B" using 2 3 SUB by (auto simp add: Field_Osum)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   182
  ultimately show ?thesis using 4 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   186
subsection \<open>The maxim among a finite set of ordinals\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   188
text \<open>The correct phrasing would be ``a maxim of ...", as \<open>\<le>o\<close> is only a preorder.\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
definition isOmax :: "'a rel set \<Rightarrow> 'a rel \<Rightarrow> bool"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   191
  where
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   192
    "isOmax  R r \<equiv> r \<in> R \<and> (\<forall>r' \<in> R. r' \<le>o r)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   194
definition omax :: "'a rel set \<Rightarrow> 'a rel"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   195
  where
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   196
    "omax R == SOME r. isOmax R r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   197
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
lemma exists_isOmax:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   199
  assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   200
  shows "\<exists> r. isOmax R r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
  have "finite R \<Longrightarrow> R \<noteq> {} \<longrightarrow> (\<forall> r \<in> R. Well_order r) \<longrightarrow> (\<exists> r. isOmax R r)"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   203
    apply(erule finite_induct) apply(simp add: isOmax_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
  proof(clarsimp)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
    fix r :: "('a \<times> 'a) set" and R assume *: "finite R" and **: "r \<notin> R"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   206
      and ***: "Well_order r" and ****: "\<forall>r\<in>R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   207
      and IH: "R \<noteq> {} \<longrightarrow> (\<exists>p. isOmax R p)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
    let ?R' = "insert r R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
    show "\<exists>p'. (isOmax ?R' p')"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
    proof(cases "R = {}")
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   211
      case True
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   212
      thus ?thesis
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   213
        by (simp add: "***" isOmax_def ordLeq_reflexive)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   214
    next
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   215
      case False
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
      then obtain p where p: "isOmax R p" using IH by auto
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   217
      hence  "Well_order p" using **** unfolding isOmax_def by simp
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   218
      then consider  "r \<le>o p" | "p \<le>o r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   219
        using *** ordLeq_total by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   220
      then show ?thesis 
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   221
      proof cases
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   222
        case 1
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   223
        then show ?thesis
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   224
         using p unfolding isOmax_def by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   225
      next
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   226
        case 2
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   227
        then show ?thesis
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   228
          by (metis "***" insert_iff isOmax_def ordLeq_reflexive ordLeq_transitive p)
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 76949
diff changeset
   229
      qed
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
    qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
  thus ?thesis using assms by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
lemma omax_isOmax:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   236
  assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   237
  shows "isOmax R (omax R)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   238
  unfolding omax_def using assms
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   239
  by(simp add: exists_isOmax someI_ex)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
lemma omax_in:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   242
  assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   243
  shows "omax R \<in> R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   244
  using assms omax_isOmax unfolding isOmax_def by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
lemma Well_order_omax:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   247
  assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   248
  shows "Well_order (omax R)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   249
  using assms omax_in by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
lemma omax_maxim:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   252
  assumes "finite R" and "\<forall> r \<in> R. Well_order r" and "r \<in> R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   253
  shows "r \<le>o omax R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   254
  using assms omax_isOmax unfolding isOmax_def by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
lemma omax_ordLeq:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   257
  assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. r \<le>o p"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   258
  shows "omax R \<le>o p"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   259
  by (meson assms omax_in ordLeq_Well_order_simp)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   260
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
lemma omax_ordLess:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   262
  assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. r <o p"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   263
  shows "omax R <o p"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   264
  by (meson assms omax_in ordLess_Well_order_simp)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
lemma omax_ordLeq_elim:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   267
  assumes "finite R" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   268
    and "omax R \<le>o p" and "r \<in> R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   269
  shows "r \<le>o p"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   270
  by (meson assms omax_maxim ordLeq_transitive)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
lemma omax_ordLess_elim:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   273
  assumes "finite R" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   274
    and "omax R <o p" and "r \<in> R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   275
  shows "r <o p"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   276
  by (meson assms omax_maxim ordLeq_ordLess_trans)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
lemma ordLeq_omax:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   279
  assumes "finite R" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   280
    and "r \<in> R" and "p \<le>o r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   281
  shows "p \<le>o omax R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   282
  by (meson assms omax_maxim ordLeq_transitive)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
lemma ordLess_omax:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   285
  assumes "finite R" and "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   286
    and "r \<in> R" and "p <o r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   287
  shows "p <o omax R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   288
  by (meson assms omax_maxim ordLess_ordLeq_trans)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
lemma omax_ordLeq_mono:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   291
  assumes P: "finite P" and R: "finite R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   292
    and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   293
    and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p \<le>o r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   294
  shows "omax P \<le>o omax R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   295
  by (meson LEQ NE_P P R Well_R omax_ordLeq ordLeq_omax)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   297
lemma omax_ordLess_mono:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   298
  assumes P: "finite P" and R: "finite R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   299
    and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   300
    and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p <o r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   301
  shows "omax P <o omax R"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   302
  by (meson LEQ NE_P P R Well_R omax_ordLess ordLess_omax)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   304
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   305
subsection \<open>Limit and succesor ordinals\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   306
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   307
lemma embed_underS2:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   308
  assumes r: "Well_order r" and g: "embed r s g" and a: "a \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   309
  shows "g ` underS r a = underS s (g a)"
72126
5de9a5fbf2ec adjustments for fewer WO assumptions
paulson <lp15@cam.ac.uk>
parents: 70078
diff changeset
   310
  by (meson a bij_betw_def embed_underS g r)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   311
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   312
lemma bij_betw_insert:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   313
  assumes "b \<notin> A" and "f b \<notin> A'" and "bij_betw f A A'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   314
  shows "bij_betw f (insert b A) (insert (f b) A')"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   315
  using notIn_Un_bij_betw[OF assms] by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   316
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   317
context wo_rel
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   318
begin
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   319
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   320
lemma underS_induct:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   321
  assumes "\<And>a. (\<And> a1. a1 \<in> underS a \<Longrightarrow> P a1) \<Longrightarrow> P a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   322
  shows "P a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   323
  by (induct rule: well_order_induct) (rule assms, simp add: underS_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   324
81808
f575ad8dcbe5 avoid theorem name clash (by Jan van Brügge)
traytel
parents: 80067
diff changeset
   325
lemma suc_underS':
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   326
  assumes B: "B \<subseteq> Field r" and A: "AboveS B \<noteq> {}" and b: "b \<in> B"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   327
  shows "b \<in> underS (suc B)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   328
  using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   329
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   330
lemma underS_supr:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   331
  assumes bA: "b \<in> underS (supr A)" and A: "A \<subseteq> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   332
  shows "\<exists> a \<in> A. b \<in> underS a"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   333
proof(rule ccontr, simp)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   334
  have bb: "b \<in> Field r" using bA unfolding underS_def Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   335
  assume "\<forall>a\<in>A.  b \<notin> underS a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   336
  hence 0: "\<forall>a \<in> A. (a,b) \<in> r" using A bA unfolding underS_def
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   337
    by simp (metis REFL in_mono max2_def max2_greater refl_on_domain)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   338
  have "(supr A, b) \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   339
    by (simp add: "0" A bb supr_least)
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   340
  thus False
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   341
    by (metis antisymD bA underS_E wo_rel.ANTISYM wo_rel_axioms)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   342
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   343
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   344
lemma underS_suc:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   345
  assumes bA: "b \<in> underS (suc A)" and A: "A \<subseteq> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   346
  shows "\<exists> a \<in> A. b \<in> under a"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   347
proof(rule ccontr, simp)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   348
  have bb: "b \<in> Field r" using bA unfolding underS_def Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   349
  assume "\<forall>a\<in>A.  b \<notin> under a"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   350
  hence 0: "\<forall>a \<in> A. a \<in> underS b" using A bA
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   351
    by (metis bb in_mono max2_def max2_greater mem_Collect_eq underS_I under_def)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   352
  have "(suc A, b) \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   353
    by (metis "0" A bb suc_least underS_E)
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   354
  thus False
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   355
    by (metis antisymD bA underS_E wo_rel.ANTISYM wo_rel_axioms)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   356
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   357
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   358
lemma (in wo_rel) in_underS_supr:
76949
ec4c38e156c7 Fixed a broken proof
paulson <lp15@cam.ac.uk>
parents: 76948
diff changeset
   359
  assumes "j \<in> underS i" and "i \<in> A" and "A \<subseteq> Field r" and "Above A \<noteq> {}"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   360
  shows "j \<in> underS (supr A)"
76949
ec4c38e156c7 Fixed a broken proof
paulson <lp15@cam.ac.uk>
parents: 76948
diff changeset
   361
  by (meson assms LIN in_mono supr_greater supr_inField underS_incl_iff)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   362
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   363
lemma inj_on_Field:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   364
  assumes A: "A \<subseteq> Field r" and f: "\<And> a b. \<lbrakk>a \<in> A; b \<in> A; a \<in> underS b\<rbrakk> \<Longrightarrow> f a \<noteq> f b"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   365
  shows "inj_on f A"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   366
  by (smt (verit) A f in_notinI inj_on_def subsetD underS_I)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   367
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   368
lemma ofilter_init_seg_of:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   369
  assumes "ofilter F"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   370
  shows "Restr r F initial_segment_of r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   371
  using assms unfolding ofilter_def init_seg_of_def under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   372
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   373
lemma underS_init_seg_of_Collect:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   374
  assumes "\<And>j1 j2. \<lbrakk>j2 \<in> underS i; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   375
  shows "{R j |j. j \<in> underS i} \<in> Chains init_seg_of"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   376
  using TOTALS assms 
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   377
  by (clarsimp simp: Chains_def) (meson BNF_Least_Fixpoint.underS_Field)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   378
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   379
lemma (in wo_rel) Field_init_seg_of_Collect:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   380
  assumes "\<And>j1 j2. \<lbrakk>j2 \<in> Field r; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   381
  shows "{R j |j. j \<in> Field r} \<in> Chains init_seg_of"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   382
  using TOTALS assms by (auto simp: Chains_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   383
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   384
subsubsection \<open>Successor and limit elements of an ordinal\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   385
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   386
definition "succ i \<equiv> suc {i}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   387
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   388
definition "isSucc i \<equiv> \<exists> j. aboveS j \<noteq> {} \<and> i = succ j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   389
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   390
definition "zero = minim (Field r)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   391
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   392
definition "isLim i \<equiv> \<not> isSucc i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   393
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   394
lemma zero_smallest[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   395
  assumes "j \<in> Field r" shows "(zero, j) \<in> r"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   396
  by (simp add: assms wo_rel.ofilter_linord wo_rel_axioms zero_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   397
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   398
lemma zero_in_Field: assumes "Field r \<noteq> {}"  shows "zero \<in> Field r"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   399
  using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   400
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   401
lemma leq_zero_imp[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   402
  "(x, zero) \<in> r \<Longrightarrow> x = zero"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   403
  by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   404
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   405
lemma leq_zero[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   406
  assumes "Field r \<noteq> {}"  shows "(x, zero) \<in> r \<longleftrightarrow> x = zero"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   407
  using zero_in_Field[OF assms] in_notinI[of x zero] by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   408
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   409
lemma under_zero[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   410
  assumes "Field r \<noteq> {}" shows "under zero = {zero}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   411
  using assms unfolding under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   412
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   413
lemma underS_zero[simp,intro]: "underS zero = {}"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   414
  unfolding underS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   415
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   416
lemma isSucc_succ: "aboveS i \<noteq> {} \<Longrightarrow> isSucc (succ i)"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   417
  unfolding isSucc_def succ_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   418
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   419
lemma succ_in_diff:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   420
  assumes "aboveS i \<noteq> {}"  shows "(i,succ i) \<in> r \<and> succ i \<noteq> i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   421
  using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   422
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   423
lemmas succ_in[simp] = succ_in_diff[THEN conjunct1]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   424
lemmas succ_diff[simp] = succ_in_diff[THEN conjunct2]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   425
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   426
lemma succ_in_Field[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   427
  assumes "aboveS i \<noteq> {}"  shows "succ i \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   428
  using succ_in[OF assms] unfolding Field_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   429
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   430
lemma succ_not_in:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   431
  assumes "aboveS i \<noteq> {}" shows "(succ i, i) \<notin> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   432
  by (metis FieldI2 assms max2_equals1 max2_equals2 succ_diff succ_in)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   433
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   434
lemma not_isSucc_zero: "\<not> isSucc zero"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   435
  by (metis isSucc_def leq_zero_imp succ_in_diff)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   436
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   437
lemma isLim_zero[simp]: "isLim zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   438
  by (metis isLim_def not_isSucc_zero)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   439
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   440
lemma succ_smallest:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   441
  assumes "(i,j) \<in> r" and "i \<noteq> j"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   442
  shows "(succ i, j) \<in> r"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   443
  by (metis Field_iff assms empty_subsetI insert_subset singletonD suc_least succ_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   444
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   445
lemma isLim_supr:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   446
  assumes f: "i \<in> Field r" and l: "isLim i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   447
  shows "i = supr (underS i)"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   448
proof(rule equals_supr)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   449
  fix j assume j: "j \<in> Field r" and 1: "\<And> j'. j' \<in> underS i \<Longrightarrow> (j',j) \<in> r"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   450
  show "(i,j) \<in> r" 
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   451
  proof(intro in_notinI[OF _ f j], safe)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   452
    assume ji: "(j,i) \<in> r" "j \<noteq> i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   453
    hence a: "aboveS j \<noteq> {}" unfolding aboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   454
    hence "i \<noteq> succ j" using l unfolding isLim_def isSucc_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   455
    moreover have "(succ j, i) \<in> r" using succ_smallest[OF ji] by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   456
    ultimately have "succ j \<in> underS i" unfolding underS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   457
    hence "(succ j, j) \<in> r" using 1 by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   458
    thus False using succ_not_in[OF a] by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   459
  qed
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   460
qed(use f underS_def Field_def in fastforce)+
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   461
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   462
definition "pred i \<equiv> SOME j. j \<in> Field r \<and> aboveS j \<noteq> {} \<and> succ j = i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   463
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   464
lemma pred_Field_succ:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   465
  assumes "isSucc i" shows "pred i \<in> Field r \<and> aboveS (pred i) \<noteq> {} \<and> succ (pred i) = i"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   466
proof-
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   467
  obtain j where j: "aboveS j \<noteq> {}" "i = succ j" 
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   468
    using assms unfolding isSucc_def by auto
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   469
  then obtain "j \<in> Field r" "j \<noteq> i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   470
    by (metis FieldI1 succ_diff succ_in)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   471
  then show ?thesis unfolding pred_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   472
    by (metis (mono_tags, lifting) j someI_ex)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   473
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   474
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   475
lemmas pred_Field[simp] = pred_Field_succ[THEN conjunct1]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   476
lemmas aboveS_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct1]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   477
lemmas succ_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct2]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   478
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   479
lemma isSucc_pred_in:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   480
  assumes "isSucc i"  shows "(pred i, i) \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   481
  by (metis assms pred_Field_succ succ_in)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   482
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   483
lemma isSucc_pred_diff:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   484
  assumes "isSucc i"  shows "pred i \<noteq> i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   485
  by (metis aboveS_pred assms succ_diff succ_pred)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   486
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   487
(* todo: pred maximal, pred injective? *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   488
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   489
lemma succ_inj[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   490
  assumes "aboveS i \<noteq> {}" and "aboveS j \<noteq> {}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   491
  shows "succ i = succ j \<longleftrightarrow> i = j"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   492
  by (metis FieldI1 assms succ_def succ_in supr_under under_underS_suc)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   493
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   494
lemma pred_succ[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   495
  assumes "aboveS j \<noteq> {}"  shows "pred (succ j) = j"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   496
  using assms isSucc_def pred_Field_succ succ_inj by blast
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   497
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   498
lemma less_succ[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   499
  assumes "aboveS i \<noteq> {}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   500
  shows "(j, succ i) \<in> r \<longleftrightarrow> (j,i) \<in> r \<or> j = succ i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   501
  by (metis FieldI1 assms in_notinI max2_equals1 max2_equals2 max2_iff succ_in succ_smallest)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   502
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   503
lemma underS_succ[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   504
  assumes "aboveS i \<noteq> {}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   505
  shows "underS (succ i) = under i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   506
  unfolding underS_def under_def by (auto simp: assms succ_not_in)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   507
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   508
lemma succ_mono:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   509
  assumes "aboveS j \<noteq> {}" and "(i,j) \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   510
  shows "(succ i, succ j) \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   511
  by (metis (full_types) assms less_succ succ_smallest)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   512
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   513
lemma under_succ[simp]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   514
  assumes "aboveS i \<noteq> {}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   515
  shows "under (succ i) = insert (succ i) (under i)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   516
  using less_succ[OF assms] unfolding under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   517
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   518
definition mergeSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   519
  where
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   520
    "mergeSL S L f i \<equiv> if isSucc i then S (pred i) (f (pred i)) else L f i"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   521
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   522
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   523
subsubsection \<open>Well-order recursion with (zero), succesor, and limit\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   524
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   525
definition worecSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   526
  where "worecSL S L \<equiv> worec (mergeSL S L)"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   527
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   528
definition "adm_woL L \<equiv> \<forall>f g i. isLim i \<and> (\<forall>j\<in>underS i. f j = g j) \<longrightarrow> L f i = L g i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   529
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   530
lemma mergeSL: "adm_woL L \<Longrightarrow>adm_wo (mergeSL S L)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   531
  unfolding adm_wo_def adm_woL_def isLim_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   532
  by (smt (verit, ccfv_threshold) isSucc_pred_diff isSucc_pred_in mergeSL_def underS_I)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   533
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   534
lemma worec_fixpoint1: "adm_wo H \<Longrightarrow> worec H i = H (worec H) i"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   535
  by (metis worec_fixpoint)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   536
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   537
lemma worecSL_isSucc:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   538
  assumes a: "adm_woL L" and i: "isSucc i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   539
  shows "worecSL S L i = S (pred i) (worecSL S L (pred i))"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   540
  by (metis a i mergeSL mergeSL_def worecSL_def worec_fixpoint)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   541
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   542
lemma worecSL_succ:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   543
  assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   544
  shows "worecSL S L (succ j) = S j (worecSL S L j)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   545
  by (simp add: a i isSucc_succ worecSL_isSucc)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   546
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   547
lemma worecSL_isLim:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   548
  assumes a: "adm_woL L" and i: "isLim i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   549
  shows "worecSL S L i = L (worecSL S L) i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   550
  by (metis a i isLim_def mergeSL mergeSL_def worecSL_def worec_fixpoint)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   551
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   552
definition worecZSL :: "'b \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   553
  where "worecZSL Z S L \<equiv> worecSL S (\<lambda> f a. if a = zero then Z else L f a)"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   554
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   555
lemma worecZSL_zero:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   556
  assumes a: "adm_woL L"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   557
  shows "worecZSL Z S L zero = Z"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   558
  by (smt (verit, best) adm_woL_def assms isLim_zero worecSL_isLim worecZSL_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   559
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   560
lemma worecZSL_succ:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   561
  assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   562
  shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   563
  unfolding worecZSL_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   564
  by (smt (verit, best) a adm_woL_def i worecSL_succ)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   565
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   566
lemma worecZSL_isLim:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   567
  assumes a: "adm_woL L" and "isLim i" and "i \<noteq> zero"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   568
  shows "worecZSL Z S L i = L (worecZSL Z S L) i"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   569
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   570
  let ?L = "\<lambda> f a. if a = zero then Z else L f a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   571
  have "worecZSL Z S L i = ?L (worecZSL Z S L) i"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   572
    unfolding worecZSL_def by (smt (verit, best) adm_woL_def assms worecSL_isLim)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   573
  also have "\<dots> = L (worecZSL Z S L) i" using assms by simp
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   574
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   575
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   576
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   577
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   578
subsubsection \<open>Well-order succ-lim induction\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   579
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   580
lemma ord_cases:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   581
  obtains j where "i = succ j" and "aboveS j \<noteq> {}"  | "isLim i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   582
  by (metis isLim_def isSucc_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   583
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   584
lemma well_order_inductSL[case_names Suc Lim]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   585
  assumes "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)"  "\<And>i. \<lbrakk>isLim i; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   586
  shows "P i"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   587
proof(induction rule: well_order_induct)
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   588
  case (1 x)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   589
  then show ?case     
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   590
    by (metis assms ord_cases succ_diff succ_in underS_E)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   591
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   592
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   593
lemma well_order_inductZSL[case_names Zero Suc Lim]:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   594
  assumes "P zero"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   595
    and  "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   596
    "\<And>i. \<lbrakk>isLim i; i \<noteq> zero; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   597
  shows "P i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   598
  by (metis assms well_order_inductSL)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   599
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   600
(* Succesor and limit ordinals *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   601
definition "isSuccOrd \<equiv> \<exists> j \<in> Field r. \<forall> i \<in> Field r. (i,j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   602
definition "isLimOrd \<equiv> \<not> isSuccOrd"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   603
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   604
lemma isLimOrd_succ:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   605
  assumes isLimOrd and "i \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   606
  shows "succ i \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   607
  using assms unfolding isLimOrd_def isSuccOrd_def
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   608
  by (metis REFL in_notinI refl_on_domain succ_smallest)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   609
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   610
lemma isLimOrd_aboveS:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   611
  assumes l: isLimOrd and i: "i \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   612
  shows "aboveS i \<noteq> {}"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   613
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   614
  obtain j where "j \<in> Field r" and "(j,i) \<notin> r"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   615
    using assms unfolding isLimOrd_def isSuccOrd_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   616
  hence "(i,j) \<in> r \<and> j \<noteq> i" by (metis i max2_def max2_greater)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   617
  thus ?thesis unfolding aboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   618
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   619
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   620
lemma succ_aboveS_isLimOrd:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   621
  assumes "\<forall> i \<in> Field r. aboveS i \<noteq> {} \<and> succ i \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   622
  shows isLimOrd
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   623
  using assms isLimOrd_def isSuccOrd_def succ_not_in by blast
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   624
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   625
lemma isLim_iff:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   626
  assumes l: "isLim i" and j: "j \<in> underS i"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   627
  shows "\<exists> k. k \<in> underS i \<and> j \<in> underS k"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   628
  by (metis Order_Relation.underS_Field empty_iff isLim_supr j l underS_empty underS_supr)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   629
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   630
end (* context wo_rel *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   631
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   632
abbreviation "zero \<equiv> wo_rel.zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   633
abbreviation "succ \<equiv> wo_rel.succ"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   634
abbreviation "pred \<equiv> wo_rel.pred"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   635
abbreviation "isSucc \<equiv> wo_rel.isSucc"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   636
abbreviation "isLim \<equiv> wo_rel.isLim"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   637
abbreviation "isLimOrd \<equiv> wo_rel.isLimOrd"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   638
abbreviation "isSuccOrd \<equiv> wo_rel.isSuccOrd"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   639
abbreviation "adm_woL \<equiv> wo_rel.adm_woL"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   640
abbreviation "worecSL \<equiv> wo_rel.worecSL"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   641
abbreviation "worecZSL \<equiv> wo_rel.worecZSL"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   642
55102
761e40ce91bc whitespace tuning
blanchet
parents: 55075
diff changeset
   643
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   644
subsection \<open>Projections of wellorders\<close>
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
   645
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   646
definition "oproj r s f \<equiv> Field s \<subseteq> f ` (Field r) \<and> compat r s f"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   647
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
   648
lemma oproj_in:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   649
  assumes "oproj r s f" and "(a,a') \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   650
  shows "(f a, f a') \<in> s"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   651
  using assms unfolding oproj_def compat_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   652
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   653
lemma oproj_Field:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   654
  assumes f: "oproj r s f" and a: "a \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   655
  shows "f a \<in> Field s"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   656
  using oproj_in[OF f] a unfolding Field_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   657
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   658
lemma oproj_Field2:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   659
  assumes f: "oproj r s f" and a: "b \<in> Field s"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   660
  shows "\<exists> a \<in> Field r. f a = b"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   661
  using assms unfolding oproj_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   662
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
   663
lemma oproj_under:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   664
  assumes f:  "oproj r s f" and a: "a \<in> under r a'"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   665
  shows "f a \<in> under s (f a')"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   666
  using oproj_in[OF f] a unfolding under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   667
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
   668
(* An ordinal is embedded in another whenever it is embedded as an order
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   669
(not necessarily as initial segment):*)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   670
theorem embedI:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   671
  assumes r: "Well_order r" and s: "Well_order s"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   672
    and f: "\<And> a. a \<in> Field r \<Longrightarrow> f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   673
  shows "\<exists> g. embed r s g"
58127
b7cab82f488e renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
blanchet
parents: 55102
diff changeset
   674
proof-
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 58889
diff changeset
   675
  interpret r: wo_rel r by unfold_locales (rule r)
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 58889
diff changeset
   676
  interpret s: wo_rel s by unfold_locales (rule s)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   677
  let ?G = "\<lambda> g a. suc s (g ` underS r a)"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61605
diff changeset
   678
  define g where "g = worec r ?G"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   679
  have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   680
      (*  *)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   681
  {fix a assume "a \<in> Field r"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   682
    hence "bij_betw g (under r a) (under s (g a)) \<and>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   683
          g a \<in> under s (f a)"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   684
    proof(induction a rule: r.underS_induct)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   685
      case (1 a)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   686
      hence a: "a \<in> Field r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   687
        and IH1a: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> inj_on g (under r a1)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   688
        and IH1b: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g ` under r a1 = under s (g a1)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   689
        and IH2: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> under s (f a1)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   690
        unfolding underS_def Field_def bij_betw_def by auto
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   691
      have fa: "f a \<in> Field s" using f[OF a] by auto
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   692
      have g: "g a = suc s (g ` underS r a)"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   693
        using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by blast
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   694
      have A0: "g ` underS r a \<subseteq> Field s"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   695
        using IH1b by (metis IH2 image_subsetI in_mono under_Field)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   696
      {fix a1 assume a1: "a1 \<in> underS r a"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   697
        from IH2[OF this] have "g a1 \<in> under s (f a1)" .
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   698
        moreover have "f a1 \<in> underS s (f a)" using f[OF a] a1 by auto
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   699
        ultimately have "g a1 \<in> underS s (f a)" by (metis s.ANTISYM s.TRANS under_underS_trans)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   700
      }
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   701
      hence fa_in: "f a \<in> AboveS s (g ` underS r a)" unfolding AboveS_def
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   702
        using fa by simp (metis (lifting, full_types) mem_Collect_eq underS_def)
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   703
      hence A: "AboveS s (g ` underS r a) \<noteq> {}" by auto
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   704
      have ga: "g a \<in> Field s" unfolding g using s.suc_inField[OF A0 A] .
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   705
      show ?case
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   706
        unfolding bij_betw_def
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   707
      proof (intro conjI)
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   708
        show "inj_on g (r.under a)"
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   709
          by (metis A IH1a IH1b a bij_betw_def g ga r s s.suc_greater subsetI wellorders_totally_ordered_aux)
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   710
        show "g ` r.under a = s.under (g a)"
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   711
          by (metis A A0 IH1a IH1b a bij_betw_def g ga r s s.suc_greater wellorders_totally_ordered_aux)
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   712
        show "g a \<in> s.under (f a)"
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   713
          by (simp add: fa_in g s.suc_least_AboveS under_def)
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   714
      qed
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   715
    qed
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   716
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   717
  thus ?thesis unfolding embed_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   718
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   719
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   720
corollary ordLeq_def2:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   721
  "r \<le>o s \<longleftrightarrow> Well_order r \<and> Well_order s \<and>
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   722
               (\<exists> f. \<forall> a \<in> Field r. f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a))"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   723
  using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s]
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   724
  unfolding ordLeq_def by fast
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   725
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   726
lemma iso_oproj:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   727
  assumes r: "Well_order r" and s: "Well_order s" and f: "iso r s f"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   728
  shows "oproj r s f"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   729
  by (metis embed_Field f iso_Field iso_iff iso_iff3 oproj_def r s)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   730
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   731
theorem oproj_embed:
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   732
  assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   733
  shows "\<exists> g. embed s r g"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   734
proof (rule embedI[OF s r, of "inv_into (Field r) f"], unfold underS_def, safe)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   735
  fix b assume "b \<in> Field s"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   736
  thus "inv_into (Field r) f b \<in> Field r" 
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   737
    using oproj_Field2[OF f] by (metis imageI inv_into_into)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   738
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   739
  fix a b assume "b \<in> Field s" "a \<noteq> b" "(a, b) \<in> s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   740
    "inv_into (Field r) f a = inv_into (Field r) f b"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   741
  with f show False
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   742
    by (meson FieldI1 in_mono inv_into_injective oproj_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   743
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   744
  fix a b assume *: "b \<in> Field s" "a \<noteq> b" "(a, b) \<in> s"
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   745
  { assume notin: "(inv_into (Field r) f a, inv_into (Field r) f b) \<notin> r"
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   746
    moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   747
    from *(3) have "a \<in> Field s" unfolding Field_def by auto
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   748
    then have "(inv_into (Field r) f b, inv_into (Field r) f a) \<in> r"
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   749
      by (meson "*"(1) notin f in_mono inv_into_into oproj_def r wo_rel.in_notinI wo_rel.intro)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   750
    ultimately have "(inv_into (Field r) f b, inv_into (Field r) f a) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   751
      using r by (auto simp: well_order_on_def linear_order_on_def total_on_def)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   752
    with f[unfolded oproj_def compat_def] *(1) \<open>a \<in> Field s\<close>
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   753
      f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"]
76946
5df58a471d9e Trying to clean up HOL/Cardinals
paulson <lp15@cam.ac.uk>
parents: 75624
diff changeset
   754
    have "(b, a) \<in> s" by (metis in_mono)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   755
    with *(2,3) s have False
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   756
      by (auto simp: well_order_on_def linear_order_on_def partial_order_on_def antisym_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   757
  } thus "(inv_into (Field r) f a, inv_into (Field r) f b) \<in> r" by blast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   758
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   759
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   760
corollary oproj_ordLeq:
72126
5de9a5fbf2ec adjustments for fewer WO assumptions
paulson <lp15@cam.ac.uk>
parents: 70078
diff changeset
   761
  assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
5de9a5fbf2ec adjustments for fewer WO assumptions
paulson <lp15@cam.ac.uk>
parents: 70078
diff changeset
   762
  shows "s \<le>o r"
76948
f33df7529fed Substantial simplification of HOL-Cardinals
paulson <lp15@cam.ac.uk>
parents: 76946
diff changeset
   763
  using f oproj_embed ordLess_iff ordLess_or_ordLeq r s by blast
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   764
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   765
end