src/HOL/Order_Relation.thy
author wenzelm
Thu, 08 Nov 2018 22:29:09 +0100
changeset 69272 15e9ed5b28fb
parent 68745 345ce5f262ea
child 70180 5beca7396282
permissions -rw-r--r--
isabelle update_cartouches -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
     1
(*  Title:      HOL/Order_Relation.thy
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
     2
    Author:     Tobias Nipkow
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
     4
*)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     5
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>Orders as Relations\<close>
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     7
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     8
theory Order_Relation
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
     9
imports Wfrec
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    10
begin
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    11
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    12
subsection \<open>Orders on a set\<close>
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    13
30198
922f944f03b2 name changes
nipkow
parents: 29859
diff changeset
    14
definition "preorder_on A r \<equiv> refl_on A r \<and> trans r"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    15
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    16
definition "partial_order_on A r \<equiv> preorder_on A r \<and> antisym r"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    17
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    18
definition "linear_order_on A r \<equiv> partial_order_on A r \<and> total_on A r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    19
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    20
definition "strict_linear_order_on A r \<equiv> trans r \<and> irrefl r \<and> total_on A r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    21
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    22
definition "well_order_on A r \<equiv> linear_order_on A r \<and> wf(r - Id)"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    23
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    24
lemmas order_on_defs =
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    25
  preorder_on_def partial_order_on_def linear_order_on_def
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    26
  strict_linear_order_on_def well_order_on_def
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    27
68745
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
    28
lemma partial_order_onD:
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
    29
  assumes "partial_order_on A r" shows "refl_on A r" and "trans r" and "antisym r"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
    30
  using assms unfolding partial_order_on_def preorder_on_def by auto
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    31
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    32
lemma preorder_on_empty[simp]: "preorder_on {} {}"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    33
  by (simp add: preorder_on_def trans_def)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    34
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    35
lemma partial_order_on_empty[simp]: "partial_order_on {} {}"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    36
  by (simp add: partial_order_on_def)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    37
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    38
lemma lnear_order_on_empty[simp]: "linear_order_on {} {}"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    39
  by (simp add: linear_order_on_def)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    40
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    41
lemma well_order_on_empty[simp]: "well_order_on {} {}"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    42
  by (simp add: well_order_on_def)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    43
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    44
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    45
lemma preorder_on_converse[simp]: "preorder_on A (r\<inverse>) = preorder_on A r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    46
  by (simp add: preorder_on_def)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    47
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    48
lemma partial_order_on_converse[simp]: "partial_order_on A (r\<inverse>) = partial_order_on A r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    49
  by (simp add: partial_order_on_def)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    50
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    51
lemma linear_order_on_converse[simp]: "linear_order_on A (r\<inverse>) = linear_order_on A r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    52
  by (simp add: linear_order_on_def)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    53
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    54
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    55
lemma strict_linear_order_on_diff_Id: "linear_order_on A r \<Longrightarrow> strict_linear_order_on A (r - Id)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    56
  by (simp add: order_on_defs trans_diff_Id)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    57
63563
0bcd79da075b prefer [simp] over [iff] as [iff] break HOL-UNITY
Andreas Lochbihler
parents: 63561
diff changeset
    58
lemma linear_order_on_singleton [simp]: "linear_order_on {x} {(x, x)}"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    59
  by (simp add: order_on_defs)
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    60
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    61
lemma linear_order_on_acyclic:
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    62
  assumes "linear_order_on A r"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    63
  shows "acyclic (r - Id)"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    64
  using strict_linear_order_on_diff_Id[OF assms]
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    65
  by (auto simp add: acyclic_irrefl strict_linear_order_on_def)
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    66
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    67
lemma linear_order_on_well_order_on:
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    68
  assumes "finite r"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    69
  shows "linear_order_on A r \<longleftrightarrow> well_order_on A r"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    70
  unfolding well_order_on_def
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    71
  using assms finite_acyclic_wf[OF _ linear_order_on_acyclic, of r] by blast
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
    72
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    73
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    74
subsection \<open>Orders on the field\<close>
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    75
30198
922f944f03b2 name changes
nipkow
parents: 29859
diff changeset
    76
abbreviation "Refl r \<equiv> refl_on (Field r) r"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    77
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    78
abbreviation "Preorder r \<equiv> preorder_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    79
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    80
abbreviation "Partial_order r \<equiv> partial_order_on (Field r) r"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    81
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    82
abbreviation "Total r \<equiv> total_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    83
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    84
abbreviation "Linear_order r \<equiv> linear_order_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    85
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    86
abbreviation "Well_order r \<equiv> well_order_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    87
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    88
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    89
lemma subset_Image_Image_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    90
  "Preorder r \<Longrightarrow> A \<subseteq> Field r \<Longrightarrow> B \<subseteq> Field r \<Longrightarrow>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    91
    r `` A \<subseteq> r `` B \<longleftrightarrow> (\<forall>a\<in>A.\<exists>b\<in>B. (b, a) \<in> r)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    92
  apply (simp add: preorder_on_def refl_on_def Image_def subset_eq)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    93
  apply (simp only: trans_def)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    94
  apply fast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    95
  done
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    96
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    97
lemma subset_Image1_Image1_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    98
  "Preorder r \<Longrightarrow> a \<in> Field r \<Longrightarrow> b \<in> Field r \<Longrightarrow> r `` {a} \<subseteq> r `` {b} \<longleftrightarrow> (b, a) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
    99
  by (simp add: subset_Image_Image_iff)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   100
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   101
lemma Refl_antisym_eq_Image1_Image1_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   102
  assumes "Refl r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   103
    and as: "antisym r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   104
    and abf: "a \<in> Field r" "b \<in> Field r"
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
   105
  shows "r `` {a} = r `` {b} \<longleftrightarrow> a = b"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   106
    (is "?lhs \<longleftrightarrow> ?rhs")
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
   107
proof
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   108
  assume ?lhs
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   109
  then have *: "\<And>x. (a, x) \<in> r \<longleftrightarrow> (b, x) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   110
    by (simp add: set_eq_iff)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   111
  have "(a, a) \<in> r" "(b, b) \<in> r" using \<open>Refl r\<close> abf by (simp_all add: refl_on_def)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   112
  then have "(a, b) \<in> r" "(b, a) \<in> r" using *[of a] *[of b] by simp_all
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   113
  then show ?rhs
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   114
    using \<open>antisym r\<close>[unfolded antisym_def] by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   115
next
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   116
  assume ?rhs
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   117
  then show ?lhs by fast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   118
qed
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   119
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   120
lemma Partial_order_eq_Image1_Image1_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   121
  "Partial_order r \<Longrightarrow> a \<in> Field r \<Longrightarrow> b \<in> Field r \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a = b"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   122
  by (auto simp: order_on_defs Refl_antisym_eq_Image1_Image1_iff)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   123
52182
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   124
lemma Total_Id_Field:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   125
  assumes "Total r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   126
    and not_Id: "\<not> r \<subseteq> Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   127
  shows "Field r = Field (r - Id)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   128
  using mono_Field[of "r - Id" r] Diff_subset[of r Id]
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   129
proof auto
52182
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   130
  fix a assume *: "a \<in> Field r"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   131
  from not_Id have "r \<noteq> {}" by fast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   132
  with not_Id obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" by auto
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   133
  then have "b \<noteq> c \<and> {b, c} \<subseteq> Field r" by (auto simp: Field_def)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   134
  with * obtain d where "d \<in> Field r" "d \<noteq> a" by auto
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   135
  with * \<open>Total r\<close> have "(a, d) \<in> r \<or> (d, a) \<in> r" by (simp add: total_on_def)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   136
  with \<open>d \<noteq> a\<close> show "a \<in> Field (r - Id)" unfolding Field_def by blast
52182
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   137
qed
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   138
68745
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   139
subsection\<open>Relations given by a predicate and the field\<close>
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   140
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   141
definition relation_of :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   142
  where "relation_of P A \<equiv> { (a, b) \<in> A \<times> A. P a b }"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   143
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   144
lemma Field_relation_of:
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   145
  assumes "refl_on A (relation_of P A)" shows "Field (relation_of P A) = A"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   146
  using assms unfolding refl_on_def Field_def by auto
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   147
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   148
lemma partial_order_on_relation_ofI:
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   149
  assumes refl: "\<And>a. a \<in> A \<Longrightarrow> P a a"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   150
    and trans: "\<And>a b c. \<lbrakk> a \<in> A; b \<in> A; c \<in> A \<rbrakk> \<Longrightarrow> P a b \<Longrightarrow> P b c \<Longrightarrow> P a c"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   151
    and antisym: "\<And>a b. \<lbrakk> a \<in> A; b \<in> A \<rbrakk> \<Longrightarrow> P a b \<Longrightarrow> P b a \<Longrightarrow> a = b"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   152
  shows "partial_order_on A (relation_of P A)"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   153
proof -
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   154
  from refl have "refl_on A (relation_of P A)"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   155
    unfolding refl_on_def relation_of_def by auto
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   156
  moreover have "trans (relation_of P A)" and "antisym (relation_of P A)"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   157
    unfolding relation_of_def
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   158
    by (auto intro: transI dest: trans, auto intro: antisymI dest: antisym)
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   159
  ultimately show ?thesis
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   160
    unfolding partial_order_on_def preorder_on_def by simp
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   161
qed
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   162
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   163
lemma Partial_order_relation_ofI:
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   164
  assumes "partial_order_on A (relation_of P A)" shows "Partial_order (relation_of P A)"
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   165
  using Field_relation_of assms partial_order_on_def preorder_on_def by fastforce
345ce5f262ea Zorn's lemma for relations defined by predicates
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   166
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   167
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   168
subsection \<open>Orders on a type\<close>
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   169
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   170
abbreviation "strict_linear_order \<equiv> strict_linear_order_on UNIV"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   171
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   172
abbreviation "linear_order \<equiv> linear_order_on UNIV"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   173
54551
4cd6deb430c3 fixed apparent copy-paste bug
blanchet
parents: 54482
diff changeset
   174
abbreviation "well_order \<equiv> well_order_on UNIV"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   175
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   176
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   177
subsection \<open>Order-like relations\<close>
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   178
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   179
text \<open>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   180
  In this subsection, we develop basic concepts and results pertaining
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   181
  to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   182
  total relations. We also further define upper and lower bounds operators.
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   183
\<close>
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   184
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   185
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   186
subsubsection \<open>Auxiliaries\<close>
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   187
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   188
lemma refl_on_domain: "refl_on A r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> a \<in> A \<and> b \<in> A"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   189
  by (auto simp add: refl_on_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   190
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   191
corollary well_order_on_domain: "well_order_on A r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> a \<in> A \<and> b \<in> A"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   192
  by (auto simp add: refl_on_domain order_on_defs)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   193
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   194
lemma well_order_on_Field: "well_order_on A r \<Longrightarrow> A = Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   195
  by (auto simp add: refl_on_def Field_def order_on_defs)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   196
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   197
lemma well_order_on_Well_order: "well_order_on A r \<Longrightarrow> A = Field r \<and> Well_order r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   198
  using well_order_on_Field [of A] by auto
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   199
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   200
lemma Total_subset_Id:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   201
  assumes "Total r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   202
    and "r \<subseteq> Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   203
  shows "r = {} \<or> (\<exists>a. r = {(a, a)})"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   204
proof -
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   205
  have "\<exists>a. r = {(a, a)}" if "r \<noteq> {}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   206
  proof -
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   207
    from that obtain a b where ab: "(a, b) \<in> r" by fast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   208
    with \<open>r \<subseteq> Id\<close> have "a = b" by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   209
    with ab have aa: "(a, a) \<in> r" by simp
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   210
    have "a = c \<and> a = d" if "(c, d) \<in> r" for c d
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   211
    proof -
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   212
      from that have "{a, c, d} \<subseteq> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   213
        using ab unfolding Field_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   214
      then have "((a, c) \<in> r \<or> (c, a) \<in> r \<or> a = c) \<and> ((a, d) \<in> r \<or> (d, a) \<in> r \<or> a = d)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   215
        using \<open>Total r\<close> unfolding total_on_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   216
      with \<open>r \<subseteq> Id\<close> show ?thesis by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   217
    qed
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   218
    then have "r \<subseteq> {(a, a)}" by auto
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   219
    with aa show ?thesis by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   220
  qed
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   221
  then show ?thesis by blast
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   222
qed
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   223
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   224
lemma Linear_order_in_diff_Id:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   225
  assumes "Linear_order r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   226
    and "a \<in> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   227
    and "b \<in> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   228
  shows "(a, b) \<in> r \<longleftrightarrow> (b, a) \<notin> r - Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   229
  using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   230
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   231
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   232
subsubsection \<open>The upper and lower bounds operators\<close>
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   233
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   234
text \<open>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   235
  Here we define upper (``above") and lower (``below") bounds operators. We
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   236
  think of \<open>r\<close> as a \<^emph>\<open>non-strict\<close> relation. The suffix \<open>S\<close> at the names of
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   237
  some operators indicates that the bounds are strict -- e.g., \<open>underS a\<close> is
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   238
  the set of all strict lower bounds of \<open>a\<close> (w.r.t. \<open>r\<close>). Capitalization of
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   239
  the first letter in the name reminds that the operator acts on sets, rather
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   240
  than on individual elements.
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   241
\<close>
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   242
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   243
definition under :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   244
  where "under r a \<equiv> {b. (b, a) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   245
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   246
definition underS :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   247
  where "underS r a \<equiv> {b. b \<noteq> a \<and> (b, a) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   248
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   249
definition Under :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   250
  where "Under r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (b, a) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   251
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   252
definition UnderS :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   253
  where "UnderS r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (b, a) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   254
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   255
definition above :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   256
  where "above r a \<equiv> {b. (a, b) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   257
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   258
definition aboveS :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   259
  where "aboveS r a \<equiv> {b. b \<noteq> a \<and> (a, b) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   260
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   261
definition Above :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   262
  where "Above r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (a, b) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   263
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   264
definition AboveS :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   265
  where "AboveS r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (a, b) \<in> r}"
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   266
55173
5556470a02b7 define ofilter outside of wo_rel
traytel
parents: 55027
diff changeset
   267
definition ofilter :: "'a rel \<Rightarrow> 'a set \<Rightarrow> bool"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   268
  where "ofilter r A \<equiv> A \<subseteq> Field r \<and> (\<forall>a \<in> A. under r a \<subseteq> A)"
55173
5556470a02b7 define ofilter outside of wo_rel
traytel
parents: 55027
diff changeset
   269
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   270
text \<open>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   271
  Note: In the definitions of \<open>Above[S]\<close> and \<open>Under[S]\<close>, we bounded
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   272
  comprehension by \<open>Field r\<close> in order to properly cover the case of \<open>A\<close> being
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   273
  empty.
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   274
\<close>
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   275
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   276
lemma underS_subset_under: "underS r a \<subseteq> under r a"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   277
  by (auto simp add: underS_def under_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   278
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   279
lemma underS_notIn: "a \<notin> underS r a"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   280
  by (simp add: underS_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   281
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   282
lemma Refl_under_in: "Refl r \<Longrightarrow> a \<in> Field r \<Longrightarrow> a \<in> under r a"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   283
  by (simp add: refl_on_def under_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   284
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   285
lemma AboveS_disjoint: "A \<inter> (AboveS r A) = {}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   286
  by (auto simp add: AboveS_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   287
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   288
lemma in_AboveS_underS: "a \<in> Field r \<Longrightarrow> a \<in> AboveS r (underS r a)"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   289
  by (auto simp add: AboveS_def underS_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   290
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   291
lemma Refl_under_underS: "Refl r \<Longrightarrow> a \<in> Field r \<Longrightarrow> under r a = underS r a \<union> {a}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   292
  unfolding under_def underS_def
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   293
  using refl_on_def[of _ r] by fastforce
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   294
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   295
lemma underS_empty: "a \<notin> Field r \<Longrightarrow> underS r a = {}"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   296
  by (auto simp: Field_def underS_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   297
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   298
lemma under_Field: "under r a \<subseteq> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   299
  by (auto simp: under_def Field_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   300
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   301
lemma underS_Field: "underS r a \<subseteq> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   302
  by (auto simp: underS_def Field_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   303
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   304
lemma underS_Field2: "a \<in> Field r \<Longrightarrow> underS r a \<subset> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   305
  using underS_notIn underS_Field by fast
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   306
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   307
lemma underS_Field3: "Field r \<noteq> {} \<Longrightarrow> underS r a \<subset> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   308
  by (cases "a \<in> Field r") (auto simp: underS_Field2 underS_empty)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   309
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   310
lemma AboveS_Field: "AboveS r A \<subseteq> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   311
  by (auto simp: AboveS_def Field_def)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   312
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   313
lemma under_incr:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   314
  assumes "trans r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   315
    and "(a, b) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   316
  shows "under r a \<subseteq> under r b"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   317
  unfolding under_def
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   318
proof auto
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   319
  fix x assume "(x, a) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   320
  with assms trans_def[of r] show "(x, b) \<in> r" by blast
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   321
qed
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   322
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   323
lemma underS_incr:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   324
  assumes "trans r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   325
    and "antisym r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   326
    and ab: "(a, b) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   327
  shows "underS r a \<subseteq> underS r b"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   328
  unfolding underS_def
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   329
proof auto
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   330
  assume *: "b \<noteq> a" and **: "(b, a) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   331
  with \<open>antisym r\<close> antisym_def[of r] ab show False
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   332
    by blast
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   333
next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   334
  fix x assume "x \<noteq> a" "(x, a) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   335
  with ab \<open>trans r\<close> trans_def[of r] show "(x, b) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   336
    by blast
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   337
qed
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   338
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   339
lemma underS_incl_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   340
  assumes LO: "Linear_order r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   341
    and INa: "a \<in> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   342
    and INb: "b \<in> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   343
  shows "underS r a \<subseteq> underS r b \<longleftrightarrow> (a, b) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   344
    (is "?lhs \<longleftrightarrow> ?rhs")
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   345
proof
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   346
  assume ?rhs
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   347
  with \<open>Linear_order r\<close> show ?lhs
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   348
    by (simp add: order_on_defs underS_incr)
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   349
next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   350
  assume *: ?lhs
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   351
  have "(a, b) \<in> r" if "a = b"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   352
    using assms that by (simp add: order_on_defs refl_on_def)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   353
  moreover have False if "a \<noteq> b" "(b, a) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   354
  proof -
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   355
    from that have "b \<in> underS r a" unfolding underS_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   356
    with * have "b \<in> underS r b" by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   357
    then show ?thesis by (simp add: underS_notIn)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   358
  qed
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   359
  ultimately show "(a,b) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   360
    using assms order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast
55026
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   361
qed
258fa7b5a621 folded 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents: 54552
diff changeset
   362
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   363
lemma finite_Linear_order_induct[consumes 3, case_names step]:
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   364
  assumes "Linear_order r"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   365
    and "x \<in> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   366
    and "finite r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   367
    and step: "\<And>x. x \<in> Field r \<Longrightarrow> (\<And>y. y \<in> aboveS r x \<Longrightarrow> P y) \<Longrightarrow> P x"
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   368
  shows "P x"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   369
  using assms(2)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   370
proof (induct rule: wf_induct[of "r\<inverse> - Id"])
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   371
  case 1
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   372
  from assms(1,3) show "wf (r\<inverse> - Id)"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   373
    using linear_order_on_well_order_on linear_order_on_converse
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   374
    unfolding well_order_on_def by blast
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   375
next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   376
  case prems: (2 x)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   377
  show ?case
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   378
    by (rule step) (use prems in \<open>auto simp: aboveS_def intro: FieldI2\<close>)
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   379
qed
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 61799
diff changeset
   380
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   381
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   382
subsection \<open>Variations on Well-Founded Relations\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   383
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   384
text \<open>
68484
59793df7f853 clarified document antiquotation @{theory};
wenzelm
parents: 63952
diff changeset
   385
  This subsection contains some variations of the results from \<^theory>\<open>HOL.Wellfounded\<close>:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   386
    \<^item> means for slightly more direct definitions by well-founded recursion;
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   387
    \<^item> variations of well-founded induction;
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   388
    \<^item> means for proving a linear order to be a well-order.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   389
\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   390
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   391
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   392
subsubsection \<open>Characterizations of well-foundedness\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   393
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   394
text \<open>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   395
  A transitive relation is well-founded iff it is ``locally'' well-founded,
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   396
  i.e., iff its restriction to the lower bounds of of any element is
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   397
  well-founded.
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   398
\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   399
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   400
lemma trans_wf_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   401
  assumes "trans r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   402
  shows "wf r \<longleftrightarrow> (\<forall>a. wf (r \<inter> (r\<inverse>``{a} \<times> r\<inverse>``{a})))"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   403
proof -
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   404
  define R where "R a = r \<inter> (r\<inverse>``{a} \<times> r\<inverse>``{a})" for a
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   405
  have "wf (R a)" if "wf r" for a
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   406
    using that R_def wf_subset[of r "R a"] by auto
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   407
  moreover
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   408
  have "wf r" if *: "\<forall>a. wf(R a)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   409
    unfolding wf_def
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   410
  proof clarify
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   411
    fix phi a
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   412
    assume **: "\<forall>a. (\<forall>b. (b, a) \<in> r \<longrightarrow> phi b) \<longrightarrow> phi a"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   413
    define chi where "chi b \<longleftrightarrow> (b, a) \<in> r \<longrightarrow> phi b" for b
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   414
    with * have "wf (R a)" by auto
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   415
    then have "(\<forall>b. (\<forall>c. (c, b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b) \<longrightarrow> (\<forall>b. chi b)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   416
      unfolding wf_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   417
    also have "\<forall>b. (\<forall>c. (c, b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   418
    proof (auto simp add: chi_def R_def)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   419
      fix b
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   420
      assume "(b, a) \<in> r" and "\<forall>c. (c, b) \<in> r \<and> (c, a) \<in> r \<longrightarrow> phi c"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   421
      then have "\<forall>c. (c, b) \<in> r \<longrightarrow> phi c"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   422
        using assms trans_def[of r] by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   423
      with ** show "phi b" by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   424
    qed
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   425
    finally have  "\<forall>b. chi b" .
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   426
    with ** chi_def show "phi a" by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   427
  qed
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   428
  ultimately show ?thesis unfolding R_def by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   429
qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   430
63952
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   431
text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   432
corollary wf_finite_segments:
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   433
  assumes "irrefl r" and "trans r" and "\<And>x. finite {y. (y, x) \<in> r}"
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   434
  shows "wf (r)"
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   435
proof (clarsimp simp: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms)
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   436
  fix a
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   437
  have "trans (r \<inter> ({x. (x, a) \<in> r} \<times> {x. (x, a) \<in> r}))"
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   438
    using assms unfolding trans_def Field_def by blast
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   439
  then show "acyclic (r \<inter> {x. (x, a) \<in> r} \<times> {x. (x, a) \<in> r})"
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   440
    using assms acyclic_def assms irrefl_def by fastforce
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   441
qed
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63572
diff changeset
   442
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 60758
diff changeset
   443
text \<open>The next lemma is a variation of \<open>wf_eq_minimal\<close> from Wellfounded,
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   444
  allowing one to assume the set included in the field.\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   445
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   446
lemma wf_eq_minimal2: "wf r \<longleftrightarrow> (\<forall>A. A \<subseteq> Field r \<and> A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a', a) \<notin> r))"
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   447
proof-
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   448
  let ?phi = "\<lambda>A. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   449
  have "wf r \<longleftrightarrow> (\<forall>A. ?phi A)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   450
    apply (auto simp: ex_in_conv [THEN sym])
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   451
     apply (erule wfE_min)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   452
      apply assumption
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   453
     apply blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   454
    apply (rule wfI_min)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   455
    apply fast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   456
    done
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   457
  also have "(\<forall>A. ?phi A) \<longleftrightarrow> (\<forall>B \<subseteq> Field r. ?phi B)"
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   458
  proof
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   459
    assume "\<forall>A. ?phi A"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   460
    then show "\<forall>B \<subseteq> Field r. ?phi B" by simp
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   461
  next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   462
    assume *: "\<forall>B \<subseteq> Field r. ?phi B"
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   463
    show "\<forall>A. ?phi A"
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   464
    proof clarify
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   465
      fix A :: "'a set"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   466
      assume **: "A \<noteq> {}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   467
      define B where "B = A \<inter> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   468
      show "\<exists>a \<in> A. \<forall>a' \<in> A. (a', a) \<notin> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   469
      proof (cases "B = {}")
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   470
        case True
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   471
        with ** obtain a where a: "a \<in> A" "a \<notin> Field r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   472
          unfolding B_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   473
        with a have "\<forall>a' \<in> A. (a',a) \<notin> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   474
          unfolding Field_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   475
        with a show ?thesis by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   476
      next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   477
        case False
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   478
        have "B \<subseteq> Field r" unfolding B_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   479
        with False * obtain a where a: "a \<in> B" "\<forall>a' \<in> B. (a', a) \<notin> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   480
          by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   481
        have "(a', a) \<notin> r" if "a' \<in> A" for a'
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   482
        proof
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   483
          assume a'a: "(a', a) \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   484
          with that have "a' \<in> B" unfolding B_def Field_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   485
          with a a'a show False by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   486
        qed
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   487
        with a show ?thesis unfolding B_def by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   488
      qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   489
    qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   490
  qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   491
  finally show ?thesis by blast
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   492
qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   493
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   494
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   495
subsubsection \<open>Characterizations of well-foundedness\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   496
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   497
text \<open>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   498
  The next lemma and its corollary enable one to prove that a linear order is
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   499
  a well-order in a way which is more standard than via well-foundedness of
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   500
  the strict version of the relation.
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   501
\<close>
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   502
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   503
lemma Linear_order_wf_diff_Id:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   504
  assumes "Linear_order r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   505
  shows "wf (r - Id) \<longleftrightarrow> (\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r))"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   506
proof (cases "r \<subseteq> Id")
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   507
  case True
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   508
  then have *: "r - Id = {}" by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   509
  have "wf (r - Id)" by (simp add: *)
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   510
  moreover have "\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   511
    if *: "A \<subseteq> Field r" and **: "A \<noteq> {}" for A
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   512
  proof -
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   513
    from \<open>Linear_order r\<close> True
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   514
    obtain a where a: "r = {} \<or> r = {(a, a)}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   515
      unfolding order_on_defs using Total_subset_Id [of r] by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   516
    with * ** have "A = {a} \<and> r = {(a, a)}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   517
      unfolding Field_def by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   518
    with a show ?thesis by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   519
  qed
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   520
  ultimately show ?thesis by blast
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   521
next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   522
  case False
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   523
  with \<open>Linear_order r\<close> have Field: "Field r = Field (r - Id)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   524
    unfolding order_on_defs using Total_Id_Field [of r] by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   525
  show ?thesis
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   526
  proof
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   527
    assume *: "wf (r - Id)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   528
    show "\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   529
    proof clarify
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   530
      fix A
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   531
      assume **: "A \<subseteq> Field r" and ***: "A \<noteq> {}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   532
      then have "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   533
        using Field * unfolding wf_eq_minimal2 by simp
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   534
      moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r \<longleftrightarrow> (a', a) \<notin> r - Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   535
        using Linear_order_in_diff_Id [OF \<open>Linear_order r\<close>] ** by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   536
      ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r" by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   537
    qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   538
  next
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   539
    assume *: "\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   540
    show "wf (r - Id)"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   541
      unfolding wf_eq_minimal2
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   542
    proof clarify
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   543
      fix A
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   544
      assume **: "A \<subseteq> Field(r - Id)" and ***: "A \<noteq> {}"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   545
      then have "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   546
        using Field * by simp
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   547
      moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r \<longleftrightarrow> (a', a) \<notin> r - Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   548
        using Linear_order_in_diff_Id [OF \<open>Linear_order r\<close>] ** mono_Field[of "r - Id" r] by blast
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   549
      ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   550
        by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   551
    qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   552
  qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   553
qed
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   554
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   555
corollary Linear_order_Well_order_iff:
63572
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   556
  "Linear_order r \<Longrightarrow>
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   557
    Well_order r \<longleftrightarrow> (\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r))"
c0cbfd2b5a45 misc tuning and modernization;
wenzelm
parents: 63563
diff changeset
   558
  unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast
55027
a74ea6d75571 folded 'Wellfounded_More_FP' into 'Wellfounded'
blanchet
parents: 55026
diff changeset
   559
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   560
end