src/HOL/BNF_Wellorder_Relation.thy
author wenzelm
Mon, 26 Jun 2023 23:20:32 +0200
changeset 78209 50c5be88ad59
parent 76950 f881fd264929
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55056
b5c94200d081 renamed '_FP' files to 'BNF_' files
blanchet
parents: 55054
diff changeset
     1
(*  Title:      HOL/BNF_Wellorder_Relation.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
55059
ef2e0fb783c6 tuned comments
blanchet
parents: 55056
diff changeset
     5
Well-order relations as needed by bounded natural functors.
48975
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
     8
section \<open>Well-Order Relations as Needed by Bounded Natural Functors\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
55056
b5c94200d081 renamed '_FP' files to 'BNF_' files
blanchet
parents: 55054
diff changeset
    10
theory BNF_Wellorder_Relation
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    11
  imports Order_Relation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    14
text\<open>In this section, we develop basic concepts and results pertaining
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
to well-order relations.  Note that we consider well-order relations
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
as {\em non-strict relations},
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    17
i.e., as containing the diagonals of their fields.\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54481
diff changeset
    19
locale wo_rel =
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54481
diff changeset
    20
  fixes r :: "'a rel"
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54481
diff changeset
    21
  assumes WELL: "Well_order r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    24
text\<open>The following context encompasses all this section. In other words,
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61799
diff changeset
    25
for the whole section, we consider a fixed well-order relation \<^term>\<open>r\<close>.\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
(* context wo_rel  *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    29
abbreviation under where "under \<equiv> Order_Relation.under r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    30
abbreviation underS where "underS \<equiv> Order_Relation.underS r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    31
abbreviation Under where "Under \<equiv> Order_Relation.Under r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    32
abbreviation UnderS where "UnderS \<equiv> Order_Relation.UnderS r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    33
abbreviation above where "above \<equiv> Order_Relation.above r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    34
abbreviation aboveS where "aboveS \<equiv> Order_Relation.aboveS r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    35
abbreviation Above where "Above \<equiv> Order_Relation.Above r"
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 55023
diff changeset
    36
abbreviation AboveS where "AboveS \<equiv> Order_Relation.AboveS r"
55173
5556470a02b7 define ofilter outside of wo_rel
traytel
parents: 55101
diff changeset
    37
abbreviation ofilter where "ofilter \<equiv> Order_Relation.ofilter r"
5556470a02b7 define ofilter outside of wo_rel
traytel
parents: 55101
diff changeset
    38
lemmas ofilter_def = Order_Relation.ofilter_def[of r]
55023
38db7814481d get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54481
diff changeset
    39
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    41
subsection \<open>Auxiliaries\<close>
48975
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 REFL: "Refl r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    44
  using WELL order_on_defs[of _ r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
lemma TRANS: "trans r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    47
  using WELL order_on_defs[of _ r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
lemma ANTISYM: "antisym r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    50
  using WELL order_on_defs[of _ r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
lemma TOTAL: "Total r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    53
  using WELL order_on_defs[of _ r] by auto
48975
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
lemma TOTALS: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    56
  using REFL TOTAL refl_on_def[of _ r] total_on_def[of _ r] by force
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
lemma LIN: "Linear_order r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    59
  using WELL well_order_on_def[of _ r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
lemma WF: "wf (r - Id)"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    62
  using WELL well_order_on_def[of _ r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
lemma cases_Total:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    65
  "\<And> phi a b. \<lbrakk>{a,b} <= Field r; ((a,b) \<in> r \<Longrightarrow> phi a b); ((b,a) \<in> r \<Longrightarrow> phi a b)\<rbrakk>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
             \<Longrightarrow> phi a b"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    67
  using TOTALS by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
lemma cases_Total3:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    70
  "\<And> phi a b. \<lbrakk>{a,b} \<le> Field r; ((a,b) \<in> r - Id \<or> (b,a) \<in> r - Id \<Longrightarrow> phi a b);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
              (a = b \<Longrightarrow> phi a b)\<rbrakk>  \<Longrightarrow> phi a b"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    72
  using TOTALS by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    75
subsection \<open>Well-founded induction and recursion adapted to non-strict well-order relations\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    77
text\<open>Here we provide induction and recursion principles specific to {\em non-strict}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
well-order relations.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
Although minor variations of those for well-founded relations, they will be useful
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
for doing away with the tediousness of
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    81
having to take out the diagonal each time in order to switch to a well-founded relation.\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
lemma well_order_induct:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    84
  assumes IND: "\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> P y \<Longrightarrow> P x"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    85
  shows "P a"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
  have "\<And>x. \<forall>y. (y, x) \<in> r - Id \<longrightarrow> P y \<Longrightarrow> P x"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    88
    using IND by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    89
  thus "P a" using WF wf_induct[of "r - Id" P a] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
definition
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    93
  worec :: "(('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    94
  where
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    95
    "worec F \<equiv> wfrec (r - Id) F"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
definition
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    98
  adm_wo :: "(('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
    99
  where
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   100
    "adm_wo H \<equiv> \<forall>f g x. (\<forall>y \<in> underS x. f y = g y) \<longrightarrow> H f x = H g x"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
lemma worec_fixpoint:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   103
  assumes ADM: "adm_wo H"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   104
  shows "worec H = H (worec H)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
  let ?rS = "r - Id"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
  have "adm_wf (r - Id) H"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   108
    unfolding adm_wf_def
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   109
    using ADM adm_wo_def[of H] underS_def[of r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
  hence "wfrec ?rS H = H (wfrec ?rS H)"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   111
    using WF wfrec_fixpoint[of ?rS H] by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
  thus ?thesis unfolding worec_def .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   116
subsection \<open>The notions of maximum, minimum, supremum, successor and order filter\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   118
text\<open>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
We define the successor {\em of a set}, and not of an element (the latter is of course
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 60758
diff changeset
   120
a particular case).  Also, we define the maximum {\em of two elements}, \<open>max2\<close>,
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 60758
diff changeset
   121
and the minimum {\em of a set}, \<open>minim\<close> -- we chose these variants since we
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
consider them the most useful for well-orders.  The minimum is defined in terms of the
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 60758
diff changeset
   123
auxiliary relational operator \<open>isMinim\<close>.  Then, supremum and successor are
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
defined in terms of minimum as expected.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
The minimum is only meaningful for non-empty sets, and the successor is only
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
meaningful for sets for which strict upper bounds exist.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   127
Order filters for well-orders are also known as ``initial segments".\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
definition max2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   130
  where "max2 a b \<equiv> if (a,b) \<in> r then b else a"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
definition isMinim :: "'a set \<Rightarrow> 'a \<Rightarrow> bool"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   133
  where "isMinim A b \<equiv> b \<in> A \<and> (\<forall>a \<in> A. (b,a) \<in> r)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
definition minim :: "'a set \<Rightarrow> 'a"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   136
  where "minim A \<equiv> THE b. isMinim A b"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
definition supr :: "'a set \<Rightarrow> 'a"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   139
  where "supr A \<equiv> minim (Above A)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
definition suc :: "'a set \<Rightarrow> 'a"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   142
  where "suc A \<equiv> minim (AboveS A)"
48975
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   145
subsubsection \<open>Properties of max2\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   147
lemma max2_greater_among:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   148
  assumes "a \<in> Field r" and "b \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   149
  shows "(a, max2 a b) \<in> r \<and> (b, max2 a b) \<in> r \<and> max2 a b \<in> {a,b}"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
  {assume "(a,b) \<in> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   152
    hence ?thesis using max2_def assms REFL refl_on_def
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   153
      by (auto simp add: refl_on_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
  {assume "a = b"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   157
    hence "(a,b) \<in> r" using REFL  assms
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   158
      by (auto simp add: refl_on_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   161
  {assume *: "a \<noteq> b \<and> (b,a) \<in> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   162
    hence "(a,b) \<notin> r" using ANTISYM
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   163
      by (auto simp add: antisym_def)
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   164
    hence ?thesis using * max2_def assms REFL refl_on_def
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   165
      by (auto simp add: refl_on_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
  ultimately show ?thesis using assms TOTAL
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   168
      total_on_def[of "Field r" r] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
qed
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 max2_greater:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   172
  assumes "a \<in> Field r" and "b \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   173
  shows "(a, max2 a b) \<in> r \<and> (b, max2 a b) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   174
  using assms by (auto simp add: max2_greater_among)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
lemma max2_among:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   177
  assumes "a \<in> Field r" and "b \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   178
  shows "max2 a b \<in> {a, b}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   179
  using assms max2_greater_among[of a b] by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
lemma max2_equals1:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   182
  assumes "a \<in> Field r" and "b \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   183
  shows "(max2 a b = a) = ((b,a) \<in> r)"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   184
  using assms ANTISYM unfolding antisym_def using TOTALS
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   185
  by(auto simp add: max2_def max2_among)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
lemma max2_equals2:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   188
  assumes "a \<in> Field r" and "b \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   189
  shows "(max2 a b = b) = ((a,b) \<in> r)"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   190
  using assms ANTISYM unfolding antisym_def using TOTALS
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   191
  unfolding max2_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   193
lemma in_notinI:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   194
  assumes "(j,i) \<notin> r \<or> j = i" and "i \<in> Field r" and "j \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   195
  shows "(i,j) \<in> r" using assms max2_def max2_greater_among by fastforce
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   197
subsubsection \<open>Existence and uniqueness for isMinim and well-definedness of minim\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
lemma isMinim_unique:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   200
  assumes "isMinim B a" "isMinim B a'"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   201
  shows "a = a'"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   202
  using assms ANTISYM antisym_def[of r] by (auto simp: isMinim_def)
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 Well_order_isMinim_exists:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   205
  assumes SUB: "B \<le> Field r" and NE: "B \<noteq> {}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   206
  shows "\<exists>b. isMinim B b"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
proof-
51764
67f05cb13e08 optimized proofs
traytel
parents: 49310
diff changeset
   208
  from spec[OF WF[unfolded wf_eq_minimal[of "r - Id"]], of B] NE obtain b where
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   209
    *: "b \<in> B \<and> (\<forall>b'. b' \<noteq> b \<and> (b',b) \<in> r \<longrightarrow> b' \<notin> B)" by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   210
  have "\<forall>b'. b' \<in> B \<longrightarrow> (b, b') \<in> r"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   211
  proof
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   212
    fix b'
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   213
    show "b' \<in> B \<longrightarrow> (b, b') \<in> r"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   214
    proof
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   215
      assume As: "b' \<in> B"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   216
      hence **: "b \<in> Field r \<and> b' \<in> Field r" using As SUB * by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   217
      from As  * have "b' = b \<or> (b',b) \<notin> r" by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   218
      moreover have "b' = b \<Longrightarrow> (b, b') \<in> r"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   219
        using ** REFL by (auto simp add: refl_on_def)
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   220
      moreover have "b' \<noteq> b \<and> (b',b) \<notin> r \<Longrightarrow> (b,b') \<in> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   221
        using ** TOTAL by (auto simp add: total_on_def)
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   222
      ultimately show "(b,b') \<in> r" by blast
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   223
    qed
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
  qed
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   225
  then show ?thesis
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   226
    unfolding isMinim_def using * by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
lemma minim_isMinim:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   230
  assumes SUB: "B \<le> Field r" and NE: "B \<noteq> {}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   231
  shows "isMinim B (minim B)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
  let ?phi = "(\<lambda> b. isMinim B b)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
  from assms Well_order_isMinim_exists
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
  obtain b where *: "?phi b" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
  have "\<And> b'. ?phi b' \<Longrightarrow> b' = b"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   238
    using isMinim_unique * by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
  ultimately show ?thesis
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   240
    unfolding minim_def using theI[of ?phi b] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   243
subsubsection\<open>Properties of minim\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
lemma minim_in:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   246
  assumes "B \<le> Field r" and "B \<noteq> {}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   247
  shows "minim B \<in> B"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   248
  using assms minim_isMinim[of B] by (auto simp: isMinim_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
lemma minim_inField:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   251
  assumes "B \<le> Field r" and "B \<noteq> {}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   252
  shows "minim B \<in> Field r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   254
  have "minim B \<in> B" using assms by (simp add: minim_in)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
  thus ?thesis using assms by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
lemma minim_least:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   259
  assumes  SUB: "B \<le> Field r" and IN: "b \<in> B"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   260
  shows "(minim B, b) \<in> r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
  from minim_isMinim[of B] assms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
  have "isMinim B (minim B)" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
  thus ?thesis by (auto simp add: isMinim_def IN)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
lemma equals_minim:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   268
  assumes SUB: "B \<le> Field r" and IN: "a \<in> B" and
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   269
    LEAST: "\<And> b. b \<in> B \<Longrightarrow> (a,b) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   270
  shows "a = minim B"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
  from minim_isMinim[of B] assms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
  have "isMinim B (minim B)" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   274
  moreover have "isMinim B a" using IN LEAST isMinim_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   275
  ultimately show ?thesis
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   276
    using isMinim_unique by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   279
subsubsection\<open>Properties of successor\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   280
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
lemma suc_AboveS:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   282
  assumes SUB: "B \<le> Field r" and ABOVES: "AboveS B \<noteq> {}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   283
  shows "suc B \<in> AboveS B"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
proof(unfold suc_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
  have "AboveS B \<le> Field r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   286
    using AboveS_Field[of r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
  thus "minim (AboveS B) \<in> AboveS B"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   288
    using assms by (simp add: minim_in)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
lemma suc_greater:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   292
  assumes SUB: "B \<le> Field r" and ABOVES: "AboveS B \<noteq> {}" and IN: "b \<in> B"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   293
  shows "suc B \<noteq> b \<and> (b,suc B) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   294
  using IN AboveS_def[of r] assms suc_AboveS by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
lemma suc_least_AboveS:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   297
  assumes ABOVES: "a \<in> AboveS B"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   298
  shows "(suc B,a) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   299
  using assms minim_least AboveS_Field[of r] by (auto simp: suc_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   301
lemma suc_inField:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   302
  assumes "B \<le> Field r" and "AboveS B \<noteq> {}"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   303
  shows "suc B \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   304
  using suc_AboveS assms AboveS_Field[of r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
lemma equals_suc_AboveS:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   307
  assumes "B \<le> Field r" and "a \<in> AboveS B" and "\<And> a'. a' \<in> AboveS B \<Longrightarrow> (a,a') \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   308
  shows "a = suc B"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   309
  using assms equals_minim AboveS_Field[of r B] by (auto simp: suc_def)
48975
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
lemma suc_underS:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   312
  assumes IN: "a \<in> Field r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   313
  shows "a = suc (underS a)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   314
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   315
  have "underS a \<le> Field r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   316
    using underS_Field[of r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   317
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   318
  have "a \<in> AboveS (underS a)"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   319
    using in_AboveS_underS IN by fast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   320
  moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
  have "\<forall>a' \<in> AboveS (underS a). (a,a') \<in> r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   322
  proof(clarify)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
    fix a'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
    assume *: "a' \<in> AboveS (underS a)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   325
    hence **: "a' \<in> Field r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   326
      using AboveS_Field by fast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
    {assume "(a,a') \<notin> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   328
      hence "a' = a \<or> (a',a) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   329
        using TOTAL IN ** by (auto simp add: total_on_def)
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   330
      moreover
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   331
      {assume "a' = a"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   332
        hence "(a,a') \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   333
          using REFL IN ** by (auto simp add: refl_on_def)
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   334
      }
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   335
      moreover
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   336
      {assume "a' \<noteq> a \<and> (a',a) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   337
        hence "a' \<in> underS a"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   338
          unfolding underS_def by simp
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   339
        hence "a' \<notin> AboveS (underS a)"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   340
          using AboveS_disjoint by fast
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   341
        with * have False by simp
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   342
      }
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   343
      ultimately have "(a,a') \<in> r" by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   344
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   345
    thus  "(a, a') \<in> r" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   346
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   347
  ultimately show ?thesis
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   348
    using equals_suc_AboveS by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   349
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   350
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   351
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   352
subsubsection \<open>Properties of order filters\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   353
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   354
lemma under_ofilter: "ofilter (under a)"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   355
  using TRANS by (auto simp: ofilter_def under_def Field_iff trans_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   356
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   357
lemma underS_ofilter: "ofilter (underS a)"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   358
  unfolding ofilter_def underS_def under_def
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   359
proof safe
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   360
  fix b assume "(a, b) \<in> r" "(b, a) \<in> r" and DIFF: "b \<noteq> a"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   361
  thus False
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   362
    using ANTISYM antisym_def[of r] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   363
next
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   364
  fix b x
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   365
  assume "(b,a) \<in> r" "b \<noteq> a" "(x,b) \<in> r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   366
  thus "(x,a) \<in> r"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   367
    using TRANS trans_def[of r] by blast
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   368
next
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   369
  fix x
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   370
  assume "x \<noteq> a" and "(x, a) \<in> r"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   371
  then show "x \<in> Field r"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   372
    unfolding Field_def
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   373
    by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   374
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   375
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   376
lemma Field_ofilter:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   377
  "ofilter (Field r)"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   378
  by(unfold ofilter_def under_def, auto simp add: Field_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   380
lemma ofilter_underS_Field:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   381
  "ofilter A = ((\<exists>a \<in> Field r. A = underS a) \<or> (A = Field r))"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
  assume "(\<exists>a\<in>Field r. A = underS a) \<or> A = Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   384
  thus "ofilter A"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   385
    by (auto simp: underS_ofilter Field_ofilter)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   386
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   387
  assume *: "ofilter A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   388
  let ?One = "(\<exists>a\<in>Field r. A = underS a)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   389
  let ?Two = "(A = Field r)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   390
  show "?One \<or> ?Two"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   391
  proof(cases ?Two)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   392
    let ?B = "(Field r) - A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   393
    let ?a = "minim ?B"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   394
    assume "A \<noteq> Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   395
    moreover have "A \<le> Field r" using * ofilter_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   396
    ultimately have 1: "?B \<noteq> {}" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   397
    hence 2: "?a \<in> Field r" using minim_inField[of ?B] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   398
    have 3: "?a \<in> ?B" using minim_in[of ?B] 1 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   399
    hence 4: "?a \<notin> A" by blast
55173
5556470a02b7 define ofilter outside of wo_rel
traytel
parents: 55101
diff changeset
   400
    have 5: "A \<le> Field r" using * ofilter_def by auto
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   401
        (*  *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   402
    moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   403
    have "A = underS ?a"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   404
    proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   405
      show "A \<le> underS ?a"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   406
      proof
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   407
        fix x assume **: "x \<in> A"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   408
        hence 11: "x \<in> Field r" using 5 by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   409
        have 12: "x \<noteq> ?a" using 4 ** by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
        have 13: "under x \<le> A" using * ofilter_def ** by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   411
        {assume "(x,?a) \<notin> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   412
          hence "(?a,x) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   413
            using TOTAL total_on_def[of "Field r" r]
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   414
              2 4 11 12 by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   415
          hence "?a \<in> under x" using under_def[of r] by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   416
          hence "?a \<in> A" using ** 13 by blast
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   417
          with 4 have False by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   418
        }
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   419
        then have "(x,?a) \<in> r" by blast
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   420
        thus "x \<in> underS ?a"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   421
          unfolding underS_def by (auto simp add: 12)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
      qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   423
    next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   424
      show "underS ?a \<le> A"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   425
      proof
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   426
        fix x
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   427
        assume **: "x \<in> underS ?a"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   428
        hence 11: "x \<in> Field r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   429
          using Field_def unfolding underS_def by fastforce
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   430
        {assume "x \<notin> A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   431
          hence "x \<in> ?B" using 11 by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
          hence "(?a,x) \<in> r" using 3 minim_least[of ?B x] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   433
          hence False
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   434
            using ANTISYM antisym_def[of r] ** unfolding underS_def by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   435
        }
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   436
        thus "x \<in> A" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   437
      qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   438
    qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   439
    ultimately have ?One using 2 by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   440
    thus ?thesis by simp
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   441
  next
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   442
    assume "A = Field r"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   443
    then show ?thesis
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75624
diff changeset
   444
      by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   445
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   446
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   447
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   448
lemma ofilter_UNION:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   449
  "(\<And> i. i \<in> I \<Longrightarrow> ofilter(A i)) \<Longrightarrow> ofilter (\<Union>i \<in> I. A i)"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   450
  unfolding ofilter_def by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   451
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   452
lemma ofilter_under_UNION:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   453
  assumes "ofilter A"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   454
  shows "A = (\<Union>a \<in> A. under a)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   455
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   456
  have "\<forall>a \<in> A. under a \<le> A"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   457
    using assms ofilter_def by auto
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58889
diff changeset
   458
  thus "(\<Union>a \<in> A. under a) \<le> A" by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   459
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   460
  have "\<forall>a \<in> A. a \<in> under a"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   461
    using REFL Refl_under_in[of r] assms ofilter_def[of A] by blast
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58889
diff changeset
   462
  thus "A \<le> (\<Union>a \<in> A. under a)" by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   463
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   464
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   465
subsubsection\<open>Other properties\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   466
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   467
lemma ofilter_linord:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   468
  assumes OF1: "ofilter A" and OF2: "ofilter B"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   469
  shows "A \<le> B \<or> B \<le> A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   470
proof(cases "A = Field r")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   471
  assume Case1: "A = Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   472
  hence "B \<le> A" using OF2 ofilter_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   473
  thus ?thesis by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   474
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   475
  assume Case2: "A \<noteq> Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   476
  with ofilter_underS_Field OF1 obtain a where
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   477
    1: "a \<in> Field r \<and> A = underS a" by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
  show ?thesis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   479
  proof(cases "B = Field r")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   480
    assume Case21: "B = Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   481
    hence "A \<le> B" using OF1 ofilter_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   482
    thus ?thesis by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   483
  next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   484
    assume Case22: "B \<noteq> Field r"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   485
    with ofilter_underS_Field OF2 obtain b where
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   486
      2: "b \<in> Field r \<and> B = underS b" by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   487
    have "a = b \<or> (a,b) \<in> r \<or> (b,a) \<in> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   488
      using 1 2 TOTAL total_on_def[of _ r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   489
    moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   490
    {assume "a = b" with 1 2 have ?thesis by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   491
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   492
    moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   493
    {assume "(a,b) \<in> r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   494
      with underS_incr[of r] TRANS ANTISYM 1 2
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   495
      have "A \<le> B" by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   496
      hence ?thesis by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   497
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   498
    moreover
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   499
    {assume "(b,a) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   500
      with underS_incr[of r] TRANS ANTISYM 1 2
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   501
      have "B \<le> A" by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   502
      hence ?thesis by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   503
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   504
    ultimately show ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   505
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   506
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   507
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   508
lemma ofilter_AboveS_Field:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   509
  assumes "ofilter A"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   510
  shows "A \<union> (AboveS A) = Field r"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   511
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   512
  show "A \<union> (AboveS A) \<le> Field r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   513
    using assms ofilter_def AboveS_Field[of r] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   514
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   515
  {fix x assume *: "x \<in> Field r" and **: "x \<notin> A"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   516
    {fix y assume ***: "y \<in> A"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   517
      with ** have 1: "y \<noteq> x" by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   518
      {assume "(y,x) \<notin> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   519
        moreover
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   520
        have "y \<in> Field r" using assms ofilter_def *** by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   521
        ultimately have "(x,y) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   522
          using 1 * TOTAL total_on_def[of _ r] by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   523
        with *** assms ofilter_def under_def[of r] have "x \<in> A" by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   524
        with ** have False by contradiction
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   525
      }
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   526
      hence "(y,x) \<in> r" by blast
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   527
      with 1 have "y \<noteq> x \<and> (y,x) \<in> r" by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   528
    }
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   529
    with * have "x \<in> AboveS A" unfolding AboveS_def by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   530
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   531
  thus "Field r \<le> A \<union> (AboveS A)" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   532
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   533
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   534
lemma suc_ofilter_in:
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   535
  assumes OF: "ofilter A" and ABOVE_NE: "AboveS A \<noteq> {}" and
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   536
    REL: "(b,suc A) \<in> r" and DIFF: "b \<noteq> suc A"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   537
  shows "b \<in> A"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   538
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   539
  have *: "suc A \<in> Field r \<and> b \<in> Field r"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   540
    using WELL REL well_order_on_domain[of "Field r"] by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   541
  {assume **: "b \<notin> A"
76950
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   542
    hence "b \<in> AboveS A"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   543
      using OF * ofilter_AboveS_Field by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   544
    hence "(suc A, b) \<in> r"
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   545
      using suc_least_AboveS by auto
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   546
    hence False using REL DIFF ANTISYM *
f881fd264929 More cleaning up proofs, plus a TeX fix
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   547
      by (auto simp add: antisym_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   548
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   549
  thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   550
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   551
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   552
end (* context wo_rel *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   553
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   554
end