# HG changeset patch # User nipkow # Date 1665507121 -7200 # Node ID 5ede2fce5b01b764cab52b57431cdba684e68317 # Parent 2f10e7a2ff01e9ba4037168be8f6e0608c08a7fb# Parent 7aa2eb860db4a44c8361ca37bc4d21451a9cbb83 merged diff -r 7aa2eb860db4 -r 5ede2fce5b01 NEWS --- a/NEWS Tue Oct 11 18:30:09 2022 +0200 +++ b/NEWS Tue Oct 11 18:52:01 2022 +0200 @@ -13,9 +13,15 @@ - Strengthened lemma total_on_singleton. Minor INCOMPATIBILITY. - Added lemmas. antisym_if_asymp + antisymp_ge[simp] + antisymp_greater[simp] antisymp_if_asymp + antisymp_le[simp] + antisymp_less[simp] irreflD irreflpD + reflp_ge[simp] + reflp_le[simp] totalp_on_singleton[simp] diff -r 7aa2eb860db4 -r 5ede2fce5b01 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Tue Oct 11 18:30:09 2022 +0200 +++ b/src/HOL/Relation.thy Tue Oct 11 18:52:01 2022 +0200 @@ -256,6 +256,12 @@ lemma reflp_mono: "reflp R \ (\x y. R x y \ Q x y) \ reflp Q" by (rule reflp_on_mono[of UNIV R Q]) simp_all +lemma (in preorder) reflp_le[simp]: "reflp (\)" + by (simp add: reflpI) + +lemma (in preorder) reflp_ge[simp]: "reflp (\)" + by (simp add: reflpI) + subsubsection \Irreflexivity\ @@ -435,6 +441,18 @@ lemma antisymp_if_asymp: "asymp R \ antisymp R" by (rule antisym_if_asym[to_pred]) +lemma (in preorder) antisymp_less[simp]: "antisymp (<)" + by (rule antisymp_if_asymp[OF asymp_less]) + +lemma (in preorder) antisymp_greater[simp]: "antisymp (>)" + by (rule antisymp_if_asymp[OF asymp_greater]) + +lemma (in order) antisymp_le[simp]: "antisymp (\)" + by (simp add: antisympI) + +lemma (in order) antisymp_ge[simp]: "antisymp (\)" + by (simp add: antisympI) + subsubsection \Transitivity\