src/HOL/Library/Order_Relation.thy
changeset 29796 33bff35f1335
parent 28952 15a4b2cf8c34
child 30198 922f944f03b2
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/Library/Order_Relation.thy	Wed Feb 11 10:51:07 2009 +0100
     1.3 @@ -0,0 +1,101 @@
     1.4 +(*  ID          : $Id$
     1.5 +    Author      : Tobias Nipkow
     1.6 +*)
     1.7 +
     1.8 +header {* Orders as Relations *}
     1.9 +
    1.10 +theory Order_Relation
    1.11 +imports Main
    1.12 +begin
    1.13 +
    1.14 +subsection{* Orders on a set *}
    1.15 +
    1.16 +definition "preorder_on A r \<equiv> refl A r \<and> trans r"
    1.17 +
    1.18 +definition "partial_order_on A r \<equiv> preorder_on A r \<and> antisym r"
    1.19 +
    1.20 +definition "linear_order_on A r \<equiv> partial_order_on A r \<and> total_on A r"
    1.21 +
    1.22 +definition "strict_linear_order_on A r \<equiv> trans r \<and> irrefl r \<and> total_on A r"
    1.23 +
    1.24 +definition "well_order_on A r \<equiv> linear_order_on A r \<and> wf(r - Id)"
    1.25 +
    1.26 +lemmas order_on_defs =
    1.27 +  preorder_on_def partial_order_on_def linear_order_on_def
    1.28 +  strict_linear_order_on_def well_order_on_def
    1.29 +
    1.30 +
    1.31 +lemma preorder_on_empty[simp]: "preorder_on {} {}"
    1.32 +by(simp add:preorder_on_def trans_def)
    1.33 +
    1.34 +lemma partial_order_on_empty[simp]: "partial_order_on {} {}"
    1.35 +by(simp add:partial_order_on_def)
    1.36 +
    1.37 +lemma lnear_order_on_empty[simp]: "linear_order_on {} {}"
    1.38 +by(simp add:linear_order_on_def)
    1.39 +
    1.40 +lemma well_order_on_empty[simp]: "well_order_on {} {}"
    1.41 +by(simp add:well_order_on_def)
    1.42 +
    1.43 +
    1.44 +lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r"
    1.45 +by (simp add:preorder_on_def)
    1.46 +
    1.47 +lemma partial_order_on_converse[simp]:
    1.48 +  "partial_order_on A (r^-1) = partial_order_on A r"
    1.49 +by (simp add: partial_order_on_def)
    1.50 +
    1.51 +lemma linear_order_on_converse[simp]:
    1.52 +  "linear_order_on A (r^-1) = linear_order_on A r"
    1.53 +by (simp add: linear_order_on_def)
    1.54 +
    1.55 +
    1.56 +lemma strict_linear_order_on_diff_Id:
    1.57 +  "linear_order_on A r \<Longrightarrow> strict_linear_order_on A (r-Id)"
    1.58 +by(simp add: order_on_defs trans_diff_Id)
    1.59 +
    1.60 +
    1.61 +subsection{* Orders on the field *}
    1.62 +
    1.63 +abbreviation "Refl r \<equiv> refl (Field r) r"
    1.64 +
    1.65 +abbreviation "Preorder r \<equiv> preorder_on (Field r) r"
    1.66 +
    1.67 +abbreviation "Partial_order r \<equiv> partial_order_on (Field r) r"
    1.68 +
    1.69 +abbreviation "Total r \<equiv> total_on (Field r) r"
    1.70 +
    1.71 +abbreviation "Linear_order r \<equiv> linear_order_on (Field r) r"
    1.72 +
    1.73 +abbreviation "Well_order r \<equiv> well_order_on (Field r) r"
    1.74 +
    1.75 +
    1.76 +lemma subset_Image_Image_iff:
    1.77 +  "\<lbrakk> Preorder r; A \<subseteq> Field r; B \<subseteq> Field r\<rbrakk> \<Longrightarrow>
    1.78 +   r `` A \<subseteq> r `` B \<longleftrightarrow> (\<forall>a\<in>A.\<exists>b\<in>B. (b,a):r)"
    1.79 +apply(auto simp add: subset_eq preorder_on_def refl_def Image_def)
    1.80 +apply metis
    1.81 +by(metis trans_def)
    1.82 +
    1.83 +lemma subset_Image1_Image1_iff:
    1.84 +  "\<lbrakk> Preorder r; a : Field r; b : Field r\<rbrakk> \<Longrightarrow> r `` {a} \<subseteq> r `` {b} \<longleftrightarrow> (b,a):r"
    1.85 +by(simp add:subset_Image_Image_iff)
    1.86 +
    1.87 +lemma Refl_antisym_eq_Image1_Image1_iff:
    1.88 +  "\<lbrakk>Refl r; antisym r; a:Field r; b:Field r\<rbrakk> \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a=b"
    1.89 +by(simp add: expand_set_eq antisym_def refl_def) metis
    1.90 +
    1.91 +lemma Partial_order_eq_Image1_Image1_iff:
    1.92 +  "\<lbrakk>Partial_order r; a:Field r; b:Field r\<rbrakk> \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a=b"
    1.93 +by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff)
    1.94 +
    1.95 +
    1.96 +subsection{* Orders on a type *}
    1.97 +
    1.98 +abbreviation "strict_linear_order \<equiv> strict_linear_order_on UNIV"
    1.99 +
   1.100 +abbreviation "linear_order \<equiv> linear_order_on UNIV"
   1.101 +
   1.102 +abbreviation "well_order r \<equiv> well_order_on UNIV"
   1.103 +
   1.104 +end