src/HOL/Order_Relation.thy
author blanchet
Thu, 16 Jan 2014 16:20:17 +0100
changeset 55017 2df6ad1dbd66
parent 54552 5d57cbec0f0f
child 55026 258fa7b5a621
permissions -rw-r--r--
adapted to move of Wfrec
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
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
     3
*)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     4
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     5
header {* Orders as Relations *}
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     6
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     7
theory Order_Relation
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
     8
imports Wellfounded
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
     9
begin
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    10
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    11
subsection{* Orders on a set *}
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    12
30198
922f944f03b2 name changes
nipkow
parents: 29859
diff changeset
    13
definition "preorder_on A r \<equiv> refl_on A r \<and> trans r"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    14
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    15
definition "partial_order_on A r \<equiv> preorder_on A r \<and> antisym r"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    16
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    17
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
    18
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    19
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
    20
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    21
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
    22
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    23
lemmas order_on_defs =
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    24
  preorder_on_def partial_order_on_def linear_order_on_def
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    25
  strict_linear_order_on_def well_order_on_def
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    26
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    27
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    28
lemma preorder_on_empty[simp]: "preorder_on {} {}"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    29
by(simp add:preorder_on_def trans_def)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    30
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    31
lemma partial_order_on_empty[simp]: "partial_order_on {} {}"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    32
by(simp add:partial_order_on_def)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    33
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    34
lemma lnear_order_on_empty[simp]: "linear_order_on {} {}"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    35
by(simp add:linear_order_on_def)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    36
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    37
lemma well_order_on_empty[simp]: "well_order_on {} {}"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    38
by(simp add:well_order_on_def)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    39
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    40
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    41
lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    42
by (simp add:preorder_on_def)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    43
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    44
lemma partial_order_on_converse[simp]:
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    45
  "partial_order_on A (r^-1) = partial_order_on A r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    46
by (simp add: partial_order_on_def)
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    47
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    48
lemma linear_order_on_converse[simp]:
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    49
  "linear_order_on A (r^-1) = linear_order_on A r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    50
by (simp add: linear_order_on_def)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    51
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    52
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    53
lemma strict_linear_order_on_diff_Id:
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    54
  "linear_order_on A r \<Longrightarrow> strict_linear_order_on A (r-Id)"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    55
by(simp add: order_on_defs trans_diff_Id)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    56
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    57
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    58
subsection{* Orders on the field *}
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    59
30198
922f944f03b2 name changes
nipkow
parents: 29859
diff changeset
    60
abbreviation "Refl r \<equiv> refl_on (Field r) r"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    61
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    62
abbreviation "Preorder r \<equiv> preorder_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    63
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    64
abbreviation "Partial_order r \<equiv> partial_order_on (Field r) r"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    65
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    66
abbreviation "Total r \<equiv> total_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    67
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    68
abbreviation "Linear_order r \<equiv> linear_order_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    69
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    70
abbreviation "Well_order r \<equiv> well_order_on (Field r) r"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    71
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    72
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    73
lemma subset_Image_Image_iff:
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    74
  "\<lbrakk> Preorder r; A \<subseteq> Field r; B \<subseteq> Field r\<rbrakk> \<Longrightarrow>
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    75
   r `` A \<subseteq> r `` B \<longleftrightarrow> (\<forall>a\<in>A.\<exists>b\<in>B. (b,a):r)"
48750
a151db85a62b tuned proofs
blanchet
parents: 39302
diff changeset
    76
unfolding preorder_on_def refl_on_def Image_def
a151db85a62b tuned proofs
blanchet
parents: 39302
diff changeset
    77
apply (simp add: subset_eq)
a151db85a62b tuned proofs
blanchet
parents: 39302
diff changeset
    78
unfolding trans_def by fast
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    79
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    80
lemma subset_Image1_Image1_iff:
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    81
  "\<lbrakk> Preorder r; a : Field r; b : Field r\<rbrakk> \<Longrightarrow> r `` {a} \<subseteq> r `` {b} \<longleftrightarrow> (b,a):r"
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    82
by(simp add:subset_Image_Image_iff)
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    83
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    84
lemma Refl_antisym_eq_Image1_Image1_iff:
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    85
  assumes r: "Refl r" and as: "antisym r" and abf: "a \<in> Field r" "b \<in> Field r"
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    86
  shows "r `` {a} = r `` {b} \<longleftrightarrow> a = b"
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    87
proof
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    88
  assume "r `` {a} = r `` {b}"
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    89
  hence e: "\<And>x. (a, x) \<in> r \<longleftrightarrow> (b, x) \<in> r" by (simp add: set_eq_iff)
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    90
  have "(a, a) \<in> r" "(b, b) \<in> r" using r abf by (simp_all add: refl_on_def)
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    91
  hence "(a, b) \<in> r" "(b, a) \<in> r" using e[of a] e[of b] by simp_all
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    92
  thus "a = b" using as[unfolded antisym_def] by blast
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
    93
qed fast
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    94
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    95
lemma Partial_order_eq_Image1_Image1_iff:
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
    96
  "\<lbrakk>Partial_order r; a:Field r; b:Field r\<rbrakk> \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a=b"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    97
by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff)
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
    98
52182
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
    99
lemma Total_Id_Field:
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   100
assumes TOT: "Total r" and NID: "\<not> (r <= Id)"
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   101
shows "Field r = Field(r - Id)"
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   102
using mono_Field[of "r - Id" r] Diff_subset[of r Id]
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   103
proof(auto)
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   104
  have "r \<noteq> {}" using NID by fast
54482
a2874c8b3558 optimized 'bad apple' method calls
blanchet
parents: 52182
diff changeset
   105
  then obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" using NID by auto
52182
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   106
  hence 1: "b \<noteq> c \<and> {b,c} \<le> Field r" by (auto simp: Field_def)
54552
5d57cbec0f0f moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents: 54551
diff changeset
   107
52182
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   108
  fix a assume *: "a \<in> Field r"
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   109
  obtain d where 2: "d \<in> Field r" and 3: "d \<noteq> a"
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   110
  using * 1 by auto
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   111
  hence "(a,d) \<in> r \<or> (d,a) \<in> r" using * TOT
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   112
  by (simp add: total_on_def)
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   113
  thus "a \<in> Field(r - Id)" using 3 unfolding Field_def by blast
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   114
qed
57b4fdc59d3b well-order extension (by Christian Sternagel)
popescua
parents: 48750
diff changeset
   115
26295
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   116
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   117
subsection{* Orders on a type *}
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   118
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   119
abbreviation "strict_linear_order \<equiv> strict_linear_order_on UNIV"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   120
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   121
abbreviation "linear_order \<equiv> linear_order_on UNIV"
afc43168ed85 More defns and thms
nipkow
parents: 26273
diff changeset
   122
54551
4cd6deb430c3 fixed apparent copy-paste bug
blanchet
parents: 54482
diff changeset
   123
abbreviation "well_order \<equiv> well_order_on UNIV"
26273
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   124
6c4d534af98d Orders as relations
nipkow
parents:
diff changeset
   125
end