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)
subsection{* Orders on a type *}
abbreviation "strict_linear_order ≡ strict_linear_order_on UNIV"
abbreviation "linear_order ≡ linear_order_on UNIV"
abbreviation "well_order r ≡ well_order_on UNIV"
end