# HG changeset patch # User blanchet # Date 1385066014 -3600 # Node ID 5d57cbec0f0ffea74d4dae5660541869e2051bdd # Parent 4cd6deb430c3168db3b84afc79110ed345268de0 moving 'Order_Relation' to 'HOL' (since it's a BNF dependency) diff -r 4cd6deb430c3 -r 5d57cbec0f0f src/HOL/Cardinals/Order_Relation_More.thy --- a/src/HOL/Cardinals/Order_Relation_More.thy Thu Nov 21 21:33:34 2013 +0100 +++ b/src/HOL/Cardinals/Order_Relation_More.thy Thu Nov 21 21:33:34 2013 +0100 @@ -8,7 +8,7 @@ header {* Basics on Order-Like Relations *} theory Order_Relation_More -imports Order_Relation_More_FP +imports Order_Relation_More_FP Main begin diff -r 4cd6deb430c3 -r 5d57cbec0f0f src/HOL/Cardinals/Order_Relation_More_FP.thy --- a/src/HOL/Cardinals/Order_Relation_More_FP.thy Thu Nov 21 21:33:34 2013 +0100 +++ b/src/HOL/Cardinals/Order_Relation_More_FP.thy Thu Nov 21 21:33:34 2013 +0100 @@ -8,7 +8,7 @@ header {* Basics on Order-Like Relations (FP) *} theory Order_Relation_More_FP -imports "~~/src/HOL/Library/Order_Relation" +imports Order_Relation begin diff -r 4cd6deb430c3 -r 5d57cbec0f0f src/HOL/Library/Order_Relation.thy --- a/src/HOL/Library/Order_Relation.thy Thu Nov 21 21:33:34 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,116 +0,0 @@ -(* Author: Tobias Nipkow *) - -header {* Orders as Relations *} - -theory Order_Relation -imports Main -begin - -subsection{* Orders on a set *} - -definition "preorder_on A r \ refl_on A r \ trans r" - -definition "partial_order_on A r \ preorder_on A r \ antisym r" - -definition "linear_order_on A r \ partial_order_on A r \ total_on A r" - -definition "strict_linear_order_on A r \ trans r \ irrefl r \ total_on A r" - -definition "well_order_on A r \ linear_order_on A r \ wf(r - Id)" - -lemmas order_on_defs = - preorder_on_def partial_order_on_def linear_order_on_def - strict_linear_order_on_def well_order_on_def - - -lemma preorder_on_empty[simp]: "preorder_on {} {}" -by(simp add:preorder_on_def trans_def) - -lemma partial_order_on_empty[simp]: "partial_order_on {} {}" -by(simp add:partial_order_on_def) - -lemma lnear_order_on_empty[simp]: "linear_order_on {} {}" -by(simp add:linear_order_on_def) - -lemma well_order_on_empty[simp]: "well_order_on {} {}" -by(simp add:well_order_on_def) - - -lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r" -by (simp add:preorder_on_def) - -lemma partial_order_on_converse[simp]: - "partial_order_on A (r^-1) = partial_order_on A r" -by (simp add: partial_order_on_def) - -lemma linear_order_on_converse[simp]: - "linear_order_on A (r^-1) = linear_order_on A r" -by (simp add: linear_order_on_def) - - -lemma strict_linear_order_on_diff_Id: - "linear_order_on A r \ strict_linear_order_on A (r-Id)" -by(simp add: order_on_defs trans_diff_Id) - - -subsection{* Orders on the field *} - -abbreviation "Refl r \ refl_on (Field r) r" - -abbreviation "Preorder r \ preorder_on (Field r) r" - -abbreviation "Partial_order r \ partial_order_on (Field r) r" - -abbreviation "Total r \ total_on (Field r) r" - -abbreviation "Linear_order r \ linear_order_on (Field r) r" - -abbreviation "Well_order r \ well_order_on (Field r) r" - - -lemma subset_Image_Image_iff: - "\ Preorder r; A \ Field r; B \ Field r\ \ - r `` A \ r `` B \ (\a\A.\b\B. (b,a):r)" -unfolding preorder_on_def refl_on_def Image_def -apply (simp add: subset_eq) -unfolding trans_def by fast - -lemma subset_Image1_Image1_iff: - "\ Preorder r; a : Field r; b : Field r\ \ r `` {a} \ r `` {b} \ (b,a):r" -by(simp add:subset_Image_Image_iff) - -lemma Refl_antisym_eq_Image1_Image1_iff: - "\Refl r; antisym r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" -by(simp add: set_eq_iff antisym_def refl_on_def) metis - -lemma Partial_order_eq_Image1_Image1_iff: - "\Partial_order r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" -by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff) - -lemma Total_Id_Field: -assumes TOT: "Total r" and NID: "\ (r <= Id)" -shows "Field r = Field(r - Id)" -using mono_Field[of "r - Id" r] Diff_subset[of r Id] -proof(auto) - have "r \ {}" using NID by fast - then obtain b and c where "b \ c \ (b,c) \ r" using NID by auto - hence 1: "b \ c \ {b,c} \ Field r" by (auto simp: Field_def) - (* *) - fix a assume *: "a \ Field r" - obtain d where 2: "d \ Field r" and 3: "d \ a" - using * 1 by auto - hence "(a,d) \ r \ (d,a) \ r" using * TOT - by (simp add: total_on_def) - thus "a \ Field(r - Id)" using 3 unfolding Field_def by blast -qed - - -subsection{* Orders on a type *} - -abbreviation "strict_linear_order \ strict_linear_order_on UNIV" - -abbreviation "linear_order \ linear_order_on UNIV" - -abbreviation "well_order \ well_order_on UNIV" - -end diff -r 4cd6deb430c3 -r 5d57cbec0f0f src/HOL/Library/Zorn.thy --- a/src/HOL/Library/Zorn.thy Thu Nov 21 21:33:34 2013 +0100 +++ b/src/HOL/Library/Zorn.thy Thu Nov 21 21:33:34 2013 +0100 @@ -10,7 +10,7 @@ header {* Zorn's Lemma *} theory Zorn -imports Order_Relation +imports Main begin subsection {* Zorn's Lemma for the Subset Relation *} diff -r 4cd6deb430c3 -r 5d57cbec0f0f src/HOL/Main.thy --- a/src/HOL/Main.thy Thu Nov 21 21:33:34 2013 +0100 +++ b/src/HOL/Main.thy Thu Nov 21 21:33:34 2013 +0100 @@ -1,7 +1,7 @@ header {* Main HOL *} theory Main -imports Predicate_Compile Nitpick Extraction Lifting_Sum List_Prefix Coinduction +imports Predicate_Compile Nitpick Extraction Lifting_Sum List_Prefix Coinduction Order_Relation begin text {* diff -r 4cd6deb430c3 -r 5d57cbec0f0f src/HOL/Order_Relation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Order_Relation.thy Thu Nov 21 21:33:34 2013 +0100 @@ -0,0 +1,125 @@ +(* Title: HOL/Order_Relation.thy + Author: Tobias Nipkow +*) + +header {* Orders as Relations *} + +theory Order_Relation +imports Wellfounded +begin + +subsection{* Orders on a set *} + +definition "preorder_on A r \ refl_on A r \ trans r" + +definition "partial_order_on A r \ preorder_on A r \ antisym r" + +definition "linear_order_on A r \ partial_order_on A r \ total_on A r" + +definition "strict_linear_order_on A r \ trans r \ irrefl r \ total_on A r" + +definition "well_order_on A r \ linear_order_on A r \ wf(r - Id)" + +lemmas order_on_defs = + preorder_on_def partial_order_on_def linear_order_on_def + strict_linear_order_on_def well_order_on_def + + +lemma preorder_on_empty[simp]: "preorder_on {} {}" +by(simp add:preorder_on_def trans_def) + +lemma partial_order_on_empty[simp]: "partial_order_on {} {}" +by(simp add:partial_order_on_def) + +lemma lnear_order_on_empty[simp]: "linear_order_on {} {}" +by(simp add:linear_order_on_def) + +lemma well_order_on_empty[simp]: "well_order_on {} {}" +by(simp add:well_order_on_def) + + +lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r" +by (simp add:preorder_on_def) + +lemma partial_order_on_converse[simp]: + "partial_order_on A (r^-1) = partial_order_on A r" +by (simp add: partial_order_on_def) + +lemma linear_order_on_converse[simp]: + "linear_order_on A (r^-1) = linear_order_on A r" +by (simp add: linear_order_on_def) + + +lemma strict_linear_order_on_diff_Id: + "linear_order_on A r \ strict_linear_order_on A (r-Id)" +by(simp add: order_on_defs trans_diff_Id) + + +subsection{* Orders on the field *} + +abbreviation "Refl r \ refl_on (Field r) r" + +abbreviation "Preorder r \ preorder_on (Field r) r" + +abbreviation "Partial_order r \ partial_order_on (Field r) r" + +abbreviation "Total r \ total_on (Field r) r" + +abbreviation "Linear_order r \ linear_order_on (Field r) r" + +abbreviation "Well_order r \ well_order_on (Field r) r" + + +lemma subset_Image_Image_iff: + "\ Preorder r; A \ Field r; B \ Field r\ \ + r `` A \ r `` B \ (\a\A.\b\B. (b,a):r)" +unfolding preorder_on_def refl_on_def Image_def +apply (simp add: subset_eq) +unfolding trans_def by fast + +lemma subset_Image1_Image1_iff: + "\ Preorder r; a : Field r; b : Field r\ \ r `` {a} \ r `` {b} \ (b,a):r" +by(simp add:subset_Image_Image_iff) + +lemma Refl_antisym_eq_Image1_Image1_iff: + assumes r: "Refl r" and as: "antisym r" and abf: "a \ Field r" "b \ Field r" + shows "r `` {a} = r `` {b} \ a = b" +proof + assume "r `` {a} = r `` {b}" + hence e: "\x. (a, x) \ r \ (b, x) \ r" by (simp add: set_eq_iff) + have "(a, a) \ r" "(b, b) \ r" using r abf by (simp_all add: refl_on_def) + hence "(a, b) \ r" "(b, a) \ r" using e[of a] e[of b] by simp_all + thus "a = b" using as[unfolded antisym_def] by blast +qed fast + +lemma Partial_order_eq_Image1_Image1_iff: + "\Partial_order r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" +by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff) + +lemma Total_Id_Field: +assumes TOT: "Total r" and NID: "\ (r <= Id)" +shows "Field r = Field(r - Id)" +using mono_Field[of "r - Id" r] Diff_subset[of r Id] +proof(auto) + have "r \ {}" using NID by fast + then obtain b and c where "b \ c \ (b,c) \ r" using NID by auto + hence 1: "b \ c \ {b,c} \ Field r" by (auto simp: Field_def) + + fix a assume *: "a \ Field r" + obtain d where 2: "d \ Field r" and 3: "d \ a" + using * 1 by auto + hence "(a,d) \ r \ (d,a) \ r" using * TOT + by (simp add: total_on_def) + thus "a \ Field(r - Id)" using 3 unfolding Field_def by blast +qed + + +subsection{* Orders on a type *} + +abbreviation "strict_linear_order \ strict_linear_order_on UNIV" + +abbreviation "linear_order \ linear_order_on UNIV" + +abbreviation "well_order \ well_order_on UNIV" + +end