src/HOL/Cardinals/Constructions_on_Wellorders.thy
author blanchet
Mon, 20 Jan 2014 18:24:55 +0100
changeset 55056 b5c94200d081
parent 55023 38db7814481d
child 55065 6d0af3c10864
permissions -rw-r--r--
renamed '_FP' files to 'BNF_' files
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 48979
diff changeset
     1
(*  Title:      HOL/Cardinals/Constructions_on_Wellorders.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
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
header {* Constructions on Wellorders *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
theory Constructions_on_Wellorders
55056
b5c94200d081 renamed '_FP' files to 'BNF_' files
blanchet
parents: 55023
diff changeset
    11
imports BNF_Constructions_on_Wellorders Wellorder_Embedding Order_Union
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
declare
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
  ordLeq_Well_order_simp[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
  not_ordLeq_iff_ordLess[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  not_ordLess_iff_ordLeq[simp]
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
    18
  Func_empty[simp]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
    19
  Func_is_emp[simp]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
    21
lemma Func_emp2[simp]: "A \<noteq> {} \<Longrightarrow> Func A {} = {}" by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
subsection {* Restriction to a set  *}
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
lemma Restr_incr2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
"r <= r' \<Longrightarrow> Restr r A <= Restr r' A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
lemma Restr_incr:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
"\<lbrakk>r \<le> r'; A \<le> A'\<rbrakk> \<Longrightarrow> Restr r A \<le> Restr r' A'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
lemma Restr_Int:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
"Restr (Restr r A) B = Restr r (A Int B)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
by blast
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
lemma Restr_iff: "(a,b) : Restr r A = (a : A \<and> b : A \<and> (a,b) : r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
by (auto simp add: Field_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
lemma Restr_subset1: "Restr r A \<le> r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
lemma Restr_subset2: "Restr r A \<le> A \<times> A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
lemma wf_Restr:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
"wf r \<Longrightarrow> wf(Restr r A)"
51764
67f05cb13e08 optimized proofs
traytel
parents: 49310
diff changeset
    48
using Restr_subset by (elim wf_subset) simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
lemma Restr_incr1:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
"A \<le> B \<Longrightarrow> Restr r A \<le> Restr r B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
by blast
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
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
subsection {* Order filters versus restrictions and embeddings  *}
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
lemma ofilter_Restr:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
assumes WELL: "Well_order r" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
        OFA: "ofilter r A" and OFB: "ofilter r B" and SUB: "A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
shows "ofilter (Restr r B) A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
  let ?rB = "Restr r B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
  hence Refl: "Refl r" by (auto simp add: wo_rel.REFL)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
  hence Field: "Field ?rB = Field r Int B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
  using Refl_Field_Restr by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
  have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
  by (auto simp add: Well_order_Restr wo_rel_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
  (* Main proof *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
  show ?thesis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
  proof(auto simp add: WellB wo_rel.ofilter_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
    fix a assume "a \<in> A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
    hence "a \<in> Field r \<and> a \<in> B" using assms Well
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
    by (auto simp add: wo_rel.ofilter_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
    with Field show "a \<in> Field(Restr r B)" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
  next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
    fix a b assume *: "a \<in> A"  and "b \<in> under (Restr r B) a"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
    hence "b \<in> under r a"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
    using WELL OFB SUB ofilter_Restr_under[of r B a] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
    thus "b \<in> A" using * Well OFA by(auto simp add: wo_rel.ofilter_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
lemma ofilter_subset_iso:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
assumes WELL: "Well_order r" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
        OFA: "ofilter r A" and OFB: "ofilter r B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
shows "(A = B) = iso (Restr r A) (Restr r B) id"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
using assms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    89
by (auto simp add: ofilter_subset_embedS_iso)
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
54476
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
    92
subsection {* Ordering the well-orders by existence of embeddings *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
corollary ordLeq_refl_on: "refl_on {r. Well_order r} ordLeq"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
using ordLeq_reflexive unfolding ordLeq_def refl_on_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
corollary ordLeq_trans: "trans ordLeq"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
using trans_def[of ordLeq] ordLeq_transitive by blast
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 ordLeq_preorder_on: "preorder_on {r. Well_order r} ordLeq"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
by(auto simp add: preorder_on_def ordLeq_refl_on ordLeq_trans)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
corollary ordIso_refl_on: "refl_on {r. Well_order r} ordIso"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
using ordIso_reflexive unfolding refl_on_def ordIso_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
by blast
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
corollary ordIso_trans: "trans ordIso"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
using trans_def[of ordIso] ordIso_transitive by blast
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
corollary ordIso_sym: "sym ordIso"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
by (auto simp add: sym_def ordIso_symmetric)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
corollary ordIso_equiv: "equiv {r. Well_order r} ordIso"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
by (auto simp add:  equiv_def ordIso_sym ordIso_refl_on ordIso_trans)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
54476
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   117
lemma ordLess_Well_order_simp[simp]:
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   118
assumes "r <o r'"
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   119
shows "Well_order r \<and> Well_order r'"
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   120
using assms unfolding ordLess_def by simp
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   121
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   122
lemma ordIso_Well_order_simp[simp]:
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   123
assumes "r =o r'"
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   124
shows "Well_order r \<and> Well_order r'"
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   125
using assms unfolding ordIso_def by simp
478b4aa26a4c moved theorems out of LFP
blanchet
parents: 54473
diff changeset
   126
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
lemma ordLess_irrefl: "irrefl ordLess"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
by(unfold irrefl_def, auto simp add: ordLess_irreflexive)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
lemma ordLess_or_ordIso:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
assumes WELL: "Well_order r" and WELL': "Well_order r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
shows "r <o r' \<or> r' <o r \<or> r =o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
unfolding ordLess_def ordIso_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
using assms embedS_or_iso[of r r'] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
corollary ordLeq_ordLess_Un_ordIso:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
"ordLeq = ordLess \<union> ordIso"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
by (auto simp add: ordLeq_iff_ordLess_or_ordIso)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
lemma not_ordLeq_ordLess:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
"r \<le>o r' \<Longrightarrow> \<not> r' <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   142
using not_ordLess_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
lemma ordIso_or_ordLess:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
assumes WELL: "Well_order r" and WELL': "Well_order r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
shows "r =o r' \<or> r <o r' \<or> r' <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   147
using assms ordLess_or_ordLeq ordLeq_iff_ordLess_or_ordIso by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   148
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   149
lemmas ord_trans = ordIso_transitive ordLeq_transitive ordLess_transitive
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
                   ordIso_ordLeq_trans ordLeq_ordIso_trans
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
                   ordIso_ordLess_trans ordLess_ordIso_trans
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
                   ordLess_ordLeq_trans ordLeq_ordLess_trans
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   153
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
lemma ofilter_ordLeq:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
assumes "Well_order r" and "ofilter r A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
shows "Restr r A \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
  have "A \<le> Field r" using assms by (auto simp add: wo_rel_def wo_rel.ofilter_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
  thus ?thesis using assms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
  by (simp add: ofilter_subset_ordLeq wo_rel.Field_ofilter
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   161
      wo_rel_def Restr_Field)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
corollary under_Restr_ordLeq:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   165
"Well_order r \<Longrightarrow> Restr r (under r a) \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
subsection {* Copy via direct images  *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
lemma Id_dir_image: "dir_image Id f \<le> Id"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
unfolding dir_image_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
lemma Un_dir_image:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
"dir_image (r1 \<union> r2) f = (dir_image r1 f) \<union> (dir_image r2 f)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
unfolding dir_image_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
lemma Int_dir_image:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
assumes "inj_on f (Field r1 \<union> Field r2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
shows "dir_image (r1 Int r2) f = (dir_image r1 f) Int (dir_image r2 f)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   182
  show "dir_image (r1 Int r2) f \<le> (dir_image r1 f) Int (dir_image r2 f)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
  using assms unfolding dir_image_def inj_on_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
  show "(dir_image r1 f) Int (dir_image r2 f) \<le> dir_image (r1 Int r2) f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
  proof(clarify)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
    fix a' b'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
    assume "(a',b') \<in> dir_image r1 f" "(a',b') \<in> dir_image r2 f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
    then obtain a1 b1 a2 b2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
    where 1: "a' = f a1 \<and> b' = f b1 \<and> a' = f a2 \<and> b' = f b2" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
          2: "(a1,b1) \<in> r1 \<and> (a2,b2) \<in> r2" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
          3: "{a1,b1} \<le> Field r1 \<and> {a2,b2} \<le> Field r2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
    unfolding dir_image_def Field_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   194
    hence "a1 = a2 \<and> b1 = b2" using assms unfolding inj_on_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   195
    hence "a' = f a1 \<and> b' = f b1 \<and> (a1,b1) \<in> r1 Int r2 \<and> (a2,b2) \<in> r1 Int r2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
    using 1 2 by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   197
    thus "(a',b') \<in> dir_image (r1 \<inter> r2) f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
    unfolding dir_image_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
52203
055c392e79cf fixed broken Cardinals and BNF due to Order_Union
popescua
parents: 52183
diff changeset
   202
(* 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
   203
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
lemma Osum_embed:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
assumes FLD: "Field r Int Field r' = {}" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
        WELL: "Well_order r" and WELL': "Well_order r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
shows "embed r (r Osum r') id"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
  have 1: "Well_order (r Osum r')"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
  using assms by (auto simp add: Osum_Well_order)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   211
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
  have "compat r (r Osum r') id"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   213
  unfolding compat_def Osum_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   214
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   215
  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
   216
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   217
  have "ofilter (r Osum r') (Field r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   218
  using 1 proof(auto simp add: wo_rel_def wo_rel.ofilter_def
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   219
                               Field_Osum under_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
    fix a b assume 2: "a \<in> Field r" and 3: "(b,a) \<in> r Osum r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   221
    moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   222
    {assume "(b,a) \<in> r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   223
     hence "a \<in> Field r'" using Field_def[of r'] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
     hence False using 2 FLD by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   225
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   226
    moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
    {assume "a \<in> Field r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
     hence False using 2 FLD by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
    ultimately
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
    show "b \<in> Field r" by (auto simp add: Osum_def Field_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
  ultimately show ?thesis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
  using assms by (auto simp add: embed_iff_compat_inj_on_ofilter)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
corollary Osum_ordLeq:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   238
assumes FLD: "Field r Int Field r' = {}" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
        WELL: "Well_order r" and WELL': "Well_order r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
shows "r \<le>o r Osum r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
using assms Osum_embed Osum_Well_order
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
unfolding ordLeq_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
lemma Well_order_embed_copy:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
assumes WELL: "well_order_on A r" and
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
        INJ: "inj_on f A" and SUB: "f ` A \<le> B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   247
shows "\<exists>r'. well_order_on B r' \<and> r \<le>o r'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
  have "bij_betw f A (f ` A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
  using INJ inj_on_imp_bij_betw by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
  then obtain r'' where "well_order_on (f ` A) r''" and 1: "r =o r''"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   252
  using WELL  Well_order_iso_copy by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
  hence 2: "Well_order r'' \<and> Field r'' = (f ` A)"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   254
  using well_order_on_Well_order 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
  let ?C = "B - (f ` A)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
  obtain r''' where "well_order_on ?C r'''"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
  using well_order_on by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
  hence 3: "Well_order r''' \<and> Field r''' = ?C"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   260
  using well_order_on_Well_order by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
  (*  *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
  let ?r' = "r'' Osum r'''"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
  have "Field r'' Int Field r''' = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
  using 2 3 by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
  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
   266
  hence 4: "r \<le>o ?r'" using 1 ordIso_ordLeq_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
  (*  *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
  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
   269
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
  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
   271
  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
   272
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   274
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   275
subsection {* The maxim among a finite set of ordinals  *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   276
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
text {* The correct phrasing would be ``a maxim of ...", as @{text "\<le>o"} is only a preorder. *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
definition isOmax :: "'a rel set \<Rightarrow> 'a rel \<Rightarrow> bool"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   280
where
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
"isOmax  R r == r \<in> R \<and> (ALL r' : R. r' \<le>o r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
definition omax :: "'a rel set \<Rightarrow> 'a rel"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
where
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
"omax R == SOME r. isOmax R r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
lemma exists_isOmax:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
shows "\<exists> r. isOmax R r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
  have "finite R \<Longrightarrow> R \<noteq> {} \<longrightarrow> (\<forall> r \<in> R. Well_order r) \<longrightarrow> (\<exists> r. isOmax R r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
  apply(erule finite_induct) apply(simp add: isOmax_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
  proof(clarsimp)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
    fix r :: "('a \<times> 'a) set" and R assume *: "finite R" and **: "r \<notin> R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
    and ***: "Well_order r" and ****: "\<forall>r\<in>R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
    and IH: "R \<noteq> {} \<longrightarrow> (\<exists>p. isOmax R p)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   297
    let ?R' = "insert r R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   298
    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
   299
    proof(cases "R = {}")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
      assume Case1: "R = {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   301
      thus ?thesis unfolding isOmax_def using ***
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   302
      by (simp add: ordLeq_reflexive)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
    next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
      assume Case2: "R \<noteq> {}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
      then obtain p where p: "isOmax R p" using IH by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
      hence 1: "Well_order p" using **** unfolding isOmax_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   307
      {assume Case21: "r \<le>o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   308
       hence "isOmax ?R' p" using p unfolding isOmax_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   309
       hence ?thesis by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   310
      }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   311
      moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   312
      {assume Case22: "p \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   313
       {fix r' assume "r' \<in> ?R'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   314
        moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   315
        {assume "r' \<in> R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
         hence "r' \<le>o p" using p unfolding isOmax_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   317
         hence "r' \<le>o r" using Case22 by(rule ordLeq_transitive)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   318
        }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   319
        moreover have "r \<le>o r" using *** by(rule ordLeq_reflexive)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   320
        ultimately have "r' \<le>o r" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
       }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   322
       hence "isOmax ?R' r" unfolding isOmax_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
       hence ?thesis by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
      }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   325
      moreover have "r \<le>o p \<or> p \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
      using 1 *** ordLeq_total by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
      ultimately show ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   328
    qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   329
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
  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
   331
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   332
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   333
lemma omax_isOmax:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   334
assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   335
shows "isOmax R (omax R)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   336
unfolding omax_def using assms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   337
by(simp add: exists_isOmax someI_ex)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   338
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   339
lemma omax_in:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   340
assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   341
shows "omax R \<in> R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   342
using assms omax_isOmax unfolding isOmax_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
lemma Well_order_omax:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   345
assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   346
shows "Well_order (omax R)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   347
using assms apply - apply(drule omax_in) by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   349
lemma omax_maxim:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   350
assumes "finite R" and "\<forall> r \<in> R. Well_order r" and "r \<in> R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   351
shows "r \<le>o omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   352
using assms omax_isOmax unfolding isOmax_def by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   353
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   354
lemma omax_ordLeq:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   355
assumes "finite R" and "R \<noteq> {}" and *: "\<forall> r \<in> R. r \<le>o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   356
shows "omax R \<le>o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   358
  have "\<forall> r \<in> R. Well_order r" using * unfolding ordLeq_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   359
  thus ?thesis using assms omax_in by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   360
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   362
lemma omax_ordLess:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   363
assumes "finite R" and "R \<noteq> {}" and *: "\<forall> r \<in> R. r <o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   364
shows "omax R <o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   365
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   366
  have "\<forall> r \<in> R. Well_order r" using * unfolding ordLess_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   367
  thus ?thesis using assms omax_in by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   368
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   369
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   370
lemma omax_ordLeq_elim:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   371
assumes "finite R" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   372
and "omax R \<le>o p" and "r \<in> R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   373
shows "r \<le>o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   374
using assms omax_maxim[of R r] apply simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   375
using ordLeq_transitive by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   376
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   377
lemma omax_ordLess_elim:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   378
assumes "finite R" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
and "omax R <o p" and "r \<in> R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   380
shows "r <o p"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   381
using assms omax_maxim[of R r] apply simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
using ordLeq_ordLess_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   384
lemma ordLeq_omax:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   385
assumes "finite R" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   386
and "r \<in> R" and "p \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   387
shows "p \<le>o omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   388
using assms omax_maxim[of R r] apply simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   389
using ordLeq_transitive by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   390
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   391
lemma ordLess_omax:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   392
assumes "finite R" and "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   393
and "r \<in> R" and "p <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   394
shows "p <o omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   395
using assms omax_maxim[of R r] apply simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   396
using ordLess_ordLeq_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   397
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   398
lemma omax_ordLeq_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   399
assumes P: "finite P" and R: "finite R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   400
and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   401
and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   402
shows "omax P \<le>o omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   403
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   404
  let ?mp = "omax P"  let ?mr = "omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   405
  {fix p assume "p : P"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   406
   then obtain r where r: "r : R" and "p \<le>o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   407
   using LEQ by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   408
   moreover have "r <=o ?mr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
   using r R Well_R omax_maxim by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
   ultimately have "p <=o ?mr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   411
   using ordLeq_transitive by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   412
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   413
  thus "?mp <=o ?mr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   414
  using NE_P P using omax_ordLeq by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   415
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   416
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   417
lemma omax_ordLess_mono:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
assumes P: "finite P" and R: "finite R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   419
and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   420
and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   421
shows "omax P <o omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   423
  let ?mp = "omax P"  let ?mr = "omax R"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   424
  {fix p assume "p : P"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   425
   then obtain r where r: "r : R" and "p <o r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   426
   using LEQ by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   427
   moreover have "r <=o ?mr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   428
   using r R Well_R omax_maxim by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   429
   ultimately have "p <o ?mr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   430
   using ordLess_ordLeq_trans by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   431
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
  thus "?mp <o ?mr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   433
  using NE_P P omax_ordLess by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   434
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   435
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
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   438
section {* Limit and Succesor Ordinals *}
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 embed_underS2:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   441
assumes r: "Well_order r" and s: "Well_order s"  and g: "embed r s g" and a: "a \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   442
shows "g ` underS r a = underS s (g a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   443
using embed_underS[OF assms] unfolding bij_betw_def by auto
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 bij_betw_insert:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   446
assumes "b \<notin> A" and "f b \<notin> A'" and "bij_betw f A A'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   447
shows "bij_betw f (insert b A) (insert (f b) A')"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   448
using notIn_Un_bij_betw[OF assms] by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   449
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   450
context wo_rel
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   451
begin
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   452
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   453
lemma underS_induct:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   454
  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
   455
  shows "P a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   456
  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
   457
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   458
lemma suc_underS:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   459
assumes B: "B \<subseteq> Field r" and A: "AboveS B \<noteq> {}" and b: "b \<in> B"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   460
shows "b \<in> underS (suc B)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   461
using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   462
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   463
lemma underS_supr:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   464
assumes bA: "b \<in> underS (supr A)" and A: "A \<subseteq> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   465
shows "\<exists> a \<in> A. b \<in> underS a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   466
proof(rule ccontr, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   467
  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
   468
  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
   469
  hence 0: "\<forall>a \<in> A. (a,b) \<in> r" using A bA unfolding underS_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   470
  by simp (metis REFL in_mono max2_def max2_greater refl_on_domain)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   471
  have "(supr A, b) \<in> r" apply(rule supr_least[OF A bb]) using 0 by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   472
  thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD)
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
lemma underS_suc:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   476
assumes bA: "b \<in> underS (suc A)" and A: "A \<subseteq> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   477
shows "\<exists> a \<in> A. b \<in> under a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   478
proof(rule ccontr, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   479
  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
   480
  assume "\<forall>a\<in>A.  b \<notin> under a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   481
  hence 0: "\<forall>a \<in> A. a \<in> underS b" using A bA unfolding underS_def
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   482
  by simp (metis (lifting) bb max2_def max2_greater mem_Collect_eq under_def set_rev_mp)
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   483
  have "(suc A, b) \<in> r" apply(rule suc_least[OF A bb]) using 0 unfolding underS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   484
  thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   485
qed
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
lemma (in wo_rel) in_underS_supr:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   488
assumes j: "j \<in> underS i" and i: "i \<in> A" and A: "A \<subseteq> Field r" and AA: "Above A \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   489
shows "j \<in> underS (supr A)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   490
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   491
  have "(i,supr A) \<in> r" using supr_greater[OF A AA i] .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   492
  thus ?thesis using j unfolding underS_def
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   493
  by simp (metis REFL TRANS max2_def max2_equals1 refl_on_domain transD)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   494
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   495
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   496
lemma inj_on_Field:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   497
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"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   498
shows "inj_on f A"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   499
unfolding inj_on_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   500
  fix a b assume a: "a \<in> A" and b: "b \<in> A" and fab: "f a = f b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   501
  {assume "a \<in> underS b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   502
   hence False using f[OF a b] fab by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   503
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   504
  moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   505
  {assume "b \<in> underS a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   506
   hence False using f[OF b a] fab by auto
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
  ultimately show "a = b" using TOTALS A a b unfolding underS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   509
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   510
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   511
lemma in_notinI:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   512
assumes "(j,i) \<notin> r \<or> j = i" and "i \<in> Field r" and "j \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   513
shows "(i,j) \<in> r" by (metis assms max2_def max2_greater_among)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   514
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   515
lemma ofilter_init_seg_of:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   516
assumes "ofilter F"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   517
shows "Restr r F initial_segment_of r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   518
using assms unfolding ofilter_def init_seg_of_def under_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   519
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   520
lemma underS_init_seg_of_Collect:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   521
assumes "\<And>j1 j2. \<lbrakk>j2 \<in> underS i; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   522
shows "{R j |j. j \<in> underS i} \<in> Chains init_seg_of"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   523
unfolding Chains_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   524
  fix j ja assume jS: "j \<in> underS i" and jaS: "ja \<in> underS i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   525
  and init: "(R ja, R j) \<notin> init_seg_of"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   526
  hence jja: "{j,ja} \<subseteq> Field r" and j: "j \<in> Field r" and ja: "ja \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   527
  and jjai: "(j,i) \<in> r" "(ja,i) \<in> r"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   528
  and i: "i \<notin> {j,ja}" unfolding Field_def underS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   529
  have jj: "(j,j) \<in> r" and jaja: "(ja,ja) \<in> r" using j ja by (metis in_notinI)+
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   530
  show "R j initial_segment_of R ja"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   531
  using jja init jjai i
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   532
  by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: underS_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   533
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   534
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   535
lemma (in wo_rel) Field_init_seg_of_Collect:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   536
assumes "\<And>j1 j2. \<lbrakk>j2 \<in> Field r; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   537
shows "{R j |j. j \<in> Field r} \<in> Chains init_seg_of"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   538
unfolding Chains_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   539
  fix j ja assume jS: "j \<in> Field r" and jaS: "ja \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   540
  and init: "(R ja, R j) \<notin> init_seg_of"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   541
  hence jja: "{j,ja} \<subseteq> Field r" and j: "j \<in> Field r" and ja: "ja \<in> Field r"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   542
  unfolding Field_def underS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   543
  have jj: "(j,j) \<in> r" and jaja: "(ja,ja) \<in> r" using j ja by (metis in_notinI)+
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   544
  show "R j initial_segment_of R ja"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   545
  using jja init
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   546
  by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: Field_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   547
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   548
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   549
subsection {* Successor and limit elements of an ordinal *}
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   550
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   551
definition "succ i \<equiv> suc {i}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   552
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   553
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
   554
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   555
definition "zero = minim (Field r)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   556
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   557
definition "isLim i \<equiv> \<not> isSucc i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   558
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   559
lemma zero_smallest[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   560
assumes "j \<in> Field r" shows "(zero, j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   561
unfolding zero_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   562
by (metis AboveS_Field assms subset_AboveS_UnderS subset_antisym subset_refl suc_def suc_least_AboveS)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   563
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   564
lemma zero_in_Field: assumes "Field r \<noteq> {}"  shows "zero \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   565
using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   566
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   567
lemma leq_zero_imp[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   568
"(x, zero) \<in> r \<Longrightarrow> x = zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   569
by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   570
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   571
lemma leq_zero[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   572
assumes "Field r \<noteq> {}"  shows "(x, zero) \<in> r \<longleftrightarrow> x = zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   573
using zero_in_Field[OF assms] in_notinI[of x zero] by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   574
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   575
lemma under_zero[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   576
assumes "Field r \<noteq> {}" shows "under zero = {zero}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   577
using assms unfolding under_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   578
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   579
lemma underS_zero[simp,intro]: "underS zero = {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   580
unfolding underS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   581
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   582
lemma isSucc_succ: "aboveS i \<noteq> {} \<Longrightarrow> isSucc (succ i)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   583
unfolding isSucc_def succ_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   584
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   585
lemma succ_in_diff:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   586
assumes "aboveS i \<noteq> {}"  shows "(i,succ i) \<in> r \<and> succ i \<noteq> i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   587
using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   588
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   589
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
   590
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
   591
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   592
lemma succ_in_Field[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   593
assumes "aboveS i \<noteq> {}"  shows "succ i \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   594
using succ_in[OF assms] unfolding Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   595
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   596
lemma succ_not_in:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   597
assumes "aboveS i \<noteq> {}" shows "(succ i, i) \<notin> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   598
proof
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   599
  assume 1: "(succ i, i) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   600
  hence "succ i \<in> Field r \<and> i \<in> Field r" unfolding Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   601
  hence "(i, succ i) \<in> r \<and> succ i \<noteq> i" using assms by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   602
  thus False using 1 by (metis ANTISYM antisymD)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   603
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   604
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   605
lemma not_isSucc_zero: "\<not> isSucc zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   606
proof
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   607
  assume "isSucc zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   608
  moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   609
  then obtain i where "aboveS i \<noteq> {}" and 1: "minim (Field r) = succ i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   610
  unfolding isSucc_def zero_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   611
  hence "succ i \<in> Field r" by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   612
  ultimately show False by (metis REFL isSucc_def minim_least refl_on_domain
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   613
    subset_refl succ_in succ_not_in zero_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   614
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   615
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   616
lemma isLim_zero[simp]: "isLim zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   617
  by (metis isLim_def not_isSucc_zero)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   618
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   619
lemma succ_smallest:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   620
assumes "(i,j) \<in> r" and "i \<noteq> j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   621
shows "(succ i, j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   622
unfolding succ_def apply(rule suc_least)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   623
using assms unfolding Field_def by auto
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_supr:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   626
assumes f: "i \<in> Field r" and l: "isLim i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   627
shows "i = supr (underS i)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   628
proof(rule equals_supr)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   629
  fix j assume j: "j \<in> Field r" and 1: "\<And> j'. j' \<in> underS i \<Longrightarrow> (j',j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   630
  show "(i,j) \<in> r" proof(intro in_notinI[OF _ f j], safe)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   631
    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
   632
    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
   633
    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
   634
    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
   635
    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
   636
    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
   637
    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
   638
  qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   639
qed(insert f, unfold underS_def Field_def, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   640
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   641
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
   642
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   643
lemma pred_Field_succ:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   644
assumes "isSucc i" shows "pred i \<in> Field r \<and> aboveS (pred i) \<noteq> {} \<and> succ (pred i) = i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   645
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   646
  obtain j where a: "aboveS j \<noteq> {}" and i: "i = succ j" using assms unfolding isSucc_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   647
  have 1: "j \<in> Field r" "j \<noteq> i" unfolding Field_def i
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   648
  using succ_diff[OF a] a unfolding aboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   649
  show ?thesis unfolding pred_def apply(rule someI_ex) using 1 i a by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   650
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   651
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   652
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
   653
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
   654
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
   655
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   656
lemma isSucc_pred_in:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   657
assumes "isSucc i"  shows "(pred i, i) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   658
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   659
  def j \<equiv> "pred i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   660
  have i: "i = succ j" using assms unfolding j_def by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   661
  have a: "aboveS j \<noteq> {}" unfolding j_def using assms by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   662
  show ?thesis unfolding j_def[symmetric] unfolding i using succ_in[OF a] .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   663
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   664
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   665
lemma isSucc_pred_diff:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   666
assumes "isSucc i"  shows "pred i \<noteq> i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   667
by (metis aboveS_pred assms succ_diff succ_pred)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   668
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   669
(* todo: pred maximal, pred injective? *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   670
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   671
lemma succ_inj[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   672
assumes "aboveS i \<noteq> {}" and "aboveS j \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   673
shows "succ i = succ j \<longleftrightarrow> i = j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   674
proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   675
  assume s: "succ i = succ j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   676
  {assume "i \<noteq> j" and "(i,j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   677
   hence "(succ i, j) \<in> r" using assms by (metis succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   678
   hence False using s assms by (metis succ_not_in)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   679
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   680
  moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   681
  {assume "i \<noteq> j" and "(j,i) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   682
   hence "(succ j, i) \<in> r" using assms by (metis succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   683
   hence False using s assms by (metis succ_not_in)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   684
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   685
  ultimately show "i = j" by (metis TOTALS WELL assms(1) assms(2) succ_in_diff well_order_on_domain)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   686
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   687
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   688
lemma pred_succ[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   689
assumes "aboveS j \<noteq> {}"  shows "pred (succ j) = j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   690
unfolding pred_def apply(rule some_equality)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   691
using assms apply(force simp: Field_def aboveS_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   692
by (metis assms succ_inj)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   693
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   694
lemma less_succ[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   695
assumes "aboveS i \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   696
shows "(j, succ i) \<in> r \<longleftrightarrow> (j,i) \<in> r \<or> j = succ i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   697
apply safe
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   698
  apply (metis WELL assms in_notinI well_order_on_domain suc_singl_pred succ_def succ_in_diff)
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   699
  apply (metis (hide_lams, full_types) REFL TRANS assms max2_def max2_equals1 refl_on_domain succ_in_Field succ_not_in transD)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   700
  apply (metis assms in_notinI succ_in_Field)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   701
done
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   702
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   703
lemma underS_succ[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   704
assumes "aboveS i \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   705
shows "underS (succ i) = under i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   706
unfolding underS_def under_def by (auto simp: assms succ_not_in)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   707
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   708
lemma succ_mono:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   709
assumes "aboveS j \<noteq> {}" and "(i,j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   710
shows "(succ i, succ j) \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   711
by (metis (full_types) assms less_succ succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   712
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   713
lemma under_succ[simp]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   714
assumes "aboveS i \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   715
shows "under (succ i) = insert (succ i) (under i)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   716
using less_succ[OF assms] unfolding under_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   717
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   718
definition mergeSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   719
where
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   720
"mergeSL S L f i \<equiv>
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   721
 if isSucc i then S (pred i) (f (pred i))
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   722
 else L f i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   723
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   724
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   725
subsection {* Well-order recursion with (zero), succesor, and limit *}
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   726
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   727
definition worecSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   728
where "worecSL S L \<equiv> worec (mergeSL S L)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   729
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   730
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
   731
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   732
lemma mergeSL:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   733
assumes "adm_woL L"  shows "adm_wo (mergeSL S L)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   734
unfolding adm_wo_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   735
  fix f g :: "'a => 'b" and i :: 'a
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   736
  assume 1: "\<forall>j\<in>underS i. f j = g j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   737
  show "mergeSL S L f i = mergeSL S L g i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   738
  proof(cases "isSucc i")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   739
    case True
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   740
    hence "pred i \<in> underS i" unfolding underS_def using isSucc_pred_in isSucc_pred_diff by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   741
    thus ?thesis using True 1 unfolding mergeSL_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   742
  next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   743
    case False hence "isLim i" unfolding isLim_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   744
    thus ?thesis using assms False 1 unfolding mergeSL_def adm_woL_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   745
  qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   746
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   747
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   748
lemma worec_fixpoint1: "adm_wo H \<Longrightarrow> worec H i = H (worec H) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   749
by (metis worec_fixpoint)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   750
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   751
lemma worecSL_isSucc:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   752
assumes a: "adm_woL L" and i: "isSucc i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   753
shows "worecSL S L i = S (pred i) (worecSL S L (pred i))"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   754
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   755
  let ?H = "mergeSL S L"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   756
  have "worecSL S L i = ?H (worec ?H) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   757
  unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   758
  also have "... = S (pred i) (worecSL S L (pred i))"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   759
  unfolding worecSL_def mergeSL_def using i by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   760
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   761
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   762
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   763
lemma worecSL_succ:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   764
assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   765
shows "worecSL S L (succ j) = S j (worecSL S L j)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   766
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   767
  def i \<equiv> "succ j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   768
  have i: "isSucc i" by (metis i i_def isSucc_succ)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   769
  have ij: "j = pred i" unfolding i_def using assms by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   770
  have 0: "succ (pred i) = i" using i by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   771
  show ?thesis unfolding ij using worecSL_isSucc[OF a i] unfolding 0 .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   772
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   773
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   774
lemma worecSL_isLim:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   775
assumes a: "adm_woL L" and i: "isLim i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   776
shows "worecSL S L i = L (worecSL S L) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   777
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   778
  let ?H = "mergeSL S L"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   779
  have "worecSL S L i = ?H (worec ?H) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   780
  unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   781
  also have "... = L (worecSL S L) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   782
  using i unfolding worecSL_def mergeSL_def isLim_def by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   783
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   784
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   785
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   786
definition worecZSL :: "'b \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   787
where "worecZSL Z S L \<equiv> worecSL S (\<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
   788
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   789
lemma worecZSL_zero:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   790
assumes a: "adm_woL L"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   791
shows "worecZSL Z S L zero = Z"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   792
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   793
  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
   794
  have "worecZSL Z S L zero = ?L (worecZSL Z S L) zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   795
  unfolding worecZSL_def apply(rule worecSL_isLim)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   796
  using assms unfolding adm_woL_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   797
  also have "... = Z" by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   798
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   799
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   800
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   801
lemma worecZSL_succ:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   802
assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   803
shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   804
unfolding worecZSL_def apply(rule  worecSL_succ)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   805
using assms unfolding adm_woL_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   806
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   807
lemma worecZSL_isLim:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   808
assumes a: "adm_woL L" and "isLim i" and "i \<noteq> zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   809
shows "worecZSL Z S L i = L (worecZSL Z S L) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   810
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   811
  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
   812
  have "worecZSL Z S L i = ?L (worecZSL Z S L) i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   813
  unfolding worecZSL_def apply(rule worecSL_isLim)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   814
  using assms unfolding adm_woL_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   815
  also have "... = L (worecZSL Z S L) i" using assms by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   816
  finally show ?thesis .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   817
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   818
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   819
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   820
subsection {* Well-order succ-lim induction: *}
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   821
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   822
lemma ord_cases:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   823
obtains j where "i = succ j" and "aboveS j \<noteq> {}"  | "isLim i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   824
by (metis isLim_def isSucc_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   825
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   826
lemma well_order_inductSL[case_names Suc Lim]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   827
assumes SUCC: "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   828
LIM: "\<And>i. \<lbrakk>isLim i; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   829
shows "P i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   830
proof(induction rule: well_order_induct)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   831
  fix i assume 0:  "\<forall>j. j \<noteq> i \<and> (j, i) \<in> r \<longrightarrow> P j"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   832
  show "P i" proof(cases i rule: ord_cases)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   833
    fix j assume i: "i = succ j" and j: "aboveS j \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   834
    hence "j \<noteq> i \<and> (j, i) \<in> r" by (metis  succ_diff succ_in)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   835
    hence 1: "P j" using 0 by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   836
    show "P i" unfolding i apply(rule SUCC) using 1 j by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   837
  qed(insert 0 LIM, unfold underS_def, auto)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   838
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   839
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   840
lemma well_order_inductZSL[case_names Zero Suc Lim]:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   841
assumes ZERO: "P zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   842
and SUCC: "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   843
LIM: "\<And>i. \<lbrakk>isLim i; i \<noteq> zero; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   844
shows "P i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   845
apply(induction rule: well_order_inductSL) using assms by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   846
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   847
(* Succesor and limit ordinals *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   848
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
   849
definition "isLimOrd \<equiv> \<not> isSuccOrd"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   850
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   851
lemma isLimOrd_succ:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   852
assumes isLimOrd and "i \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   853
shows "succ i \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   854
using assms unfolding isLimOrd_def isSuccOrd_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   855
by (metis REFL in_notinI refl_on_domain succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   856
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   857
lemma isLimOrd_aboveS:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   858
assumes l: isLimOrd and i: "i \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   859
shows "aboveS i \<noteq> {}"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   860
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   861
  obtain j where "j \<in> Field r" and "(j,i) \<notin> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   862
  using assms unfolding isLimOrd_def isSuccOrd_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   863
  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
   864
  thus ?thesis unfolding aboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   865
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   866
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   867
lemma succ_aboveS_isLimOrd:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   868
assumes "\<forall> i \<in> Field r. aboveS i \<noteq> {} \<and> succ i \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   869
shows isLimOrd
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   870
unfolding isLimOrd_def isSuccOrd_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   871
  fix j assume j: "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
   872
  hence "(succ j, j) \<in> r" using assms by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   873
  moreover have "aboveS j \<noteq> {}" using assms j unfolding aboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   874
  ultimately show False by (metis succ_not_in)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   875
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   876
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   877
lemma isLim_iff:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   878
assumes l: "isLim i" and j: "j \<in> underS i"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   879
shows "\<exists> k. k \<in> underS i \<and> j \<in> underS k"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   880
proof-
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   881
  have a: "aboveS j \<noteq> {}" using j unfolding underS_def aboveS_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   882
  show ?thesis apply(rule exI[of _ "succ j"]) apply safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   883
  using assms a unfolding underS_def isLim_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   884
  apply (metis (lifting, full_types) isSucc_def mem_Collect_eq succ_smallest)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   885
  by (metis (lifting, full_types) a mem_Collect_eq succ_diff succ_in)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   886
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   887
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   888
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   889
end (* context wo_rel *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   890
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   891
abbreviation "zero \<equiv> wo_rel.zero"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   892
abbreviation "succ \<equiv> wo_rel.succ"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   893
abbreviation "pred \<equiv> wo_rel.pred"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   894
abbreviation "isSucc \<equiv> wo_rel.isSucc"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   895
abbreviation "isLim \<equiv> wo_rel.isLim"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   896
abbreviation "isLimOrd \<equiv> wo_rel.isLimOrd"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   897
abbreviation "isSuccOrd \<equiv> wo_rel.isSuccOrd"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   898
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
   899
abbreviation "worecSL \<equiv> wo_rel.worecSL"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   900
abbreviation "worecZSL \<equiv> wo_rel.worecZSL"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   901
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   902
section {* Projections of Wellorders *}
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   903
 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   904
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
   905
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   906
lemma oproj_in: 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   907
assumes "oproj r s f" and "(a,a') \<in> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   908
shows "(f a, f a') \<in> s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   909
using assms unfolding oproj_def compat_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   910
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   911
lemma oproj_Field:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   912
assumes f: "oproj r s f" and a: "a \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   913
shows "f a \<in> Field s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   914
using oproj_in[OF f] a unfolding Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   915
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   916
lemma oproj_Field2:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   917
assumes f: "oproj r s f" and a: "b \<in> Field s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   918
shows "\<exists> a \<in> Field r. f a = b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   919
using assms unfolding oproj_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   920
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   921
lemma oproj_under: 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   922
assumes f:  "oproj r s f" and a: "a \<in> under r a'"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   923
shows "f a \<in> under s (f a')"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   924
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
   925
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   926
(* An ordinal is embedded in another whenever it is embedded as an order 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   927
(not necessarily as initial segment):*)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   928
theorem embedI:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   929
assumes r: "Well_order r" and s: "Well_order s" 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   930
and f: "\<And> a. a \<in> Field r \<Longrightarrow> f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   931
shows "\<exists> g. embed r s g"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   932
proof-  
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   933
  interpret r!: wo_rel r by unfold_locales (rule r)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   934
  interpret s!: wo_rel s by unfold_locales (rule s)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   935
  let ?G = "\<lambda> g a. suc s (g ` underS r a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   936
  def g \<equiv> "worec r ?G"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   937
  have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   938
  (*  *)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   939
  {fix a assume "a \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   940
   hence "bij_betw g (under r a) (under s (g a)) \<and> 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   941
          g a \<in> under s (f a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   942
   proof(induction a rule: r.underS_induct)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   943
     case (1 a)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   944
     hence a: "a \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   945
     and IH1a: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> inj_on g (under r a1)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   946
     and IH1b: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g ` under r a1 = under s (g a1)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   947
     and IH2: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> under s (f a1)"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   948
     unfolding underS_def Field_def bij_betw_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   949
     have fa: "f a \<in> Field s" using f[OF a] by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   950
     have g: "g a = suc s (g ` underS r a)" 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   951
     using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by simp
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   952
     have A0: "g ` underS r a \<subseteq> Field s" 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   953
     using IH1b by (metis IH2 image_subsetI in_mono under_Field)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   954
     {fix a1 assume a1: "a1 \<in> underS r a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   955
      from IH2[OF this] have "g a1 \<in> under s (f a1)" .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   956
      moreover have "f a1 \<in> underS s (f a)" using f[OF a] a1 by auto
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   957
      ultimately have "g a1 \<in> underS s (f a)" by (metis s.ANTISYM s.TRANS under_underS_trans)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   958
     }
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   959
     hence "f a \<in> AboveS s (g ` underS r a)" unfolding AboveS_def 
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   960
     using fa by simp (metis (lifting, full_types) mem_Collect_eq underS_def)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   961
     hence A: "AboveS s (g ` underS r a) \<noteq> {}" by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   962
     have B: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (g a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   963
     unfolding g apply(rule s.suc_underS[OF A0 A]) by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   964
     {fix a1 a2 assume a2: "a2 \<in> underS r a" and 1: "a1 \<in> underS r a2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   965
      hence a12: "{a1,a2} \<subseteq> under r a2" and "a1 \<noteq> a2" using r.REFL a
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   966
      unfolding underS_def under_def refl_on_def Field_def by auto 
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   967
      hence "g a1 \<noteq> g a2" using IH1a[OF a2] unfolding inj_on_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   968
      hence "g a1 \<in> underS s (g a2)" using IH1b[OF a2] a12 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   969
      unfolding underS_def under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   970
     } note C = this
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   971
     have ga: "g a \<in> Field s" unfolding g using s.suc_inField[OF A0 A] .
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   972
     have aa: "a \<in> under r a" 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   973
     using a r.REFL unfolding under_def underS_def refl_on_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   974
     show ?case proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   975
       show "bij_betw g (under r a) (under s (g a))" unfolding bij_betw_def proof safe
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   976
         show "inj_on g (under r a)" proof(rule r.inj_on_Field)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   977
           fix a1 a2 assume "a1 \<in> under r a" and a2: "a2 \<in> under r a" and a1: "a1 \<in> underS r a2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   978
           hence a22: "a2 \<in> under r a2" and a12: "a1 \<in> under r a2" "a1 \<noteq> a2"
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   979
           using a r.REFL unfolding under_def underS_def refl_on_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   980
           show "g a1 \<noteq> g a2"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   981
           proof(cases "a2 = a")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   982
             case False hence "a2 \<in> underS r a" 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   983
             using a2 unfolding underS_def under_def by auto 
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   984
             from IH1a[OF this] show ?thesis using a12 a22 unfolding inj_on_def by auto
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   985
           qed(insert B a1, unfold underS_def, auto)
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   986
         qed(unfold under_def Field_def, auto)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   987
       next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   988
         fix a1 assume a1: "a1 \<in> under r a" 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   989
         show "g a1 \<in> under s (g a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   990
         proof(cases "a1 = a")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   991
           case True thus ?thesis 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   992
           using ga s.REFL unfolding refl_on_def under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   993
         next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   994
           case False
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   995
           hence a1: "a1 \<in> underS r a" using a1 unfolding underS_def under_def by auto 
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
   996
           thus ?thesis using B unfolding underS_def under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   997
         qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   998
       next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
   999
         fix b1 assume b1: "b1 \<in> under s (g a)"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1000
         show "b1 \<in> g ` under r a"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1001
         proof(cases "b1 = g a")
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1002
           case True thus ?thesis using aa by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1003
         next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1004
           case False 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
  1005
           hence "b1 \<in> underS s (g a)" using b1 unfolding underS_def under_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1006
           from s.underS_suc[OF this[unfolded g] A0]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1007
           obtain a1 where a1: "a1 \<in> underS r a" and b1: "b1 \<in> under s (g a1)" by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1008
           obtain a2 where "a2 \<in> under r a1" and b1: "b1 = g a2" using IH1b[OF a1] b1 by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1009
           hence "a2 \<in> under r a" using a1 
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
  1010
           by (metis r.ANTISYM r.TRANS in_mono underS_subset_under under_underS_trans)
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1011
           thus ?thesis using b1 by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1012
         qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1013
       qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1014
     next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1015
       have "(g a, f a) \<in> s" unfolding g proof(rule s.suc_least[OF A0])
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1016
         fix b1 assume "b1 \<in> g ` underS r a" 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1017
         then obtain a1 where a1: "b1 = g a1" and a1: "a1 \<in> underS r a" by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1018
         hence "b1 \<in> underS s (f a)" 
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1019
         using a by (metis `\<And>a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (f a)`)
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
  1020
         thus "f a \<noteq> b1 \<and> (b1, f a) \<in> s" unfolding underS_def by auto
54980
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1021
       qed(insert fa, auto)
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54980
diff changeset
  1022
       thus "g a \<in> under s (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
  1023
     qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1024
   qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1025
  }
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1026
  thus ?thesis unfolding embed_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1027
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1028
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1029
corollary ordLeq_def2:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1030
  "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
  1031
               (\<exists> f. \<forall> a \<in> Field r. f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a))"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1032
using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1033
unfolding ordLeq_def by fast
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1034
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1035
lemma iso_oproj:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1036
  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
  1037
  shows "oproj r s f"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1038
using assms unfolding oproj_def
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1039
by (subst (asm) iso_iff3) (auto simp: bij_betw_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1040
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1041
theorem oproj_embed:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1042
assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1043
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
  1044
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
  1045
  fix b assume "b \<in> Field s"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1046
  thus "inv_into (Field r) f b \<in> Field r" using oproj_Field2[OF f] by (metis imageI inv_into_into)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1047
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1048
  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
  1049
    "inv_into (Field r) f a = inv_into (Field r) f b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1050
  with f show False by (auto dest!: inv_into_injective simp: Field_def oproj_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1051
next
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1052
  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
  1053
  { assume "(inv_into (Field r) f a, inv_into (Field r) f b) \<notin> r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1054
    moreover
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1055
    from *(3) have "a \<in> Field s" unfolding Field_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1056
    with *(1,2) have
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1057
      "inv_into (Field r) f a \<in> Field r" "inv_into (Field r) f b \<in> Field r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1058
      "inv_into (Field r) f a \<noteq> inv_into (Field r) f b"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1059
      by (auto dest!: oproj_Field2[OF f] inv_into_injective intro!: inv_into_into)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1060
    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
  1061
      using r by (auto simp: well_order_on_def linear_order_on_def total_on_def)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1062
    with f[unfolded oproj_def compat_def] *(1) `a \<in> Field s`
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1063
      f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"]
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1064
      have "(b, a) \<in> s" by (metis in_mono)
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1065
    with *(2,3) s have False
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1066
      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
  1067
  } 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
  1068
qed
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1069
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1070
corollary oproj_ordLeq:
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1071
assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1072
shows "s \<le>o r"
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1073
using oproj_embed[OF assms] r s unfolding ordLeq_def by auto
7e0573a490ee basic ordinal arithmetic and cardinals library extension (not relevant for BNFs)
traytel
parents: 54545
diff changeset
  1074
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1075
end