diff -r 5929b172c6fe -r 56d5bb8c102e src/HOL/Library/Preorder.thy --- a/src/HOL/Library/Preorder.thy Fri Feb 15 07:11:11 2019 +0000 +++ b/src/HOL/Library/Preorder.thy Fri Feb 15 18:24:22 2019 +0000 @@ -16,13 +16,22 @@ equiv ("'(\')") and equiv ("(_/ \ _)" [51, 51] 50) -lemma refl [iff]: "x \ x" +lemma equivD1: "x \ y" if "x \ y" + using that by (simp add: equiv_def) + +lemma equivD2: "y \ x" if "x \ y" + using that by (simp add: equiv_def) + +lemma equiv_refl [iff]: "x \ x" by (simp add: equiv_def) -lemma trans: "x \ y \ y \ z \ x \ z" +lemma equiv_sym: "x \ y \ y \ x" + by (auto simp add: equiv_def) + +lemma equiv_trans: "x \ y \ y \ z \ x \ z" by (auto simp: equiv_def intro: order_trans) -lemma antisym: "x \ y \ y \ x \ x \ y" +lemma equiv_antisym: "x \ y \ y \ x \ x \ y" by (simp only: equiv_def) lemma less_le: "x < y \ x \ y \ \ x \ y" @@ -31,24 +40,55 @@ lemma le_less: "x \ y \ x < y \ x \ y" by (auto simp add: equiv_def less_le) -lemma le_imp_less_or_eq: "x \ y \ x < y \ x \ y" +lemma le_imp_less_or_equiv: "x \ y \ x < y \ x \ y" by (simp add: less_le) -lemma less_imp_not_eq: "x < y \ x \ y \ False" +lemma less_imp_not_equiv: "x < y \ \ x \ y" by (simp add: less_le) -lemma less_imp_not_eq2: "x < y \ y \ x \ False" - by (simp add: equiv_def less_le) - -lemma neq_le_trans: "\ a \ b \ a \ b \ a < b" +lemma not_equiv_le_trans: "\ a \ b \ a \ b \ a < b" by (simp add: less_le) -lemma le_neq_trans: "a \ b \ \ a \ b \ a < b" - by (simp add: less_le) +lemma le_not_equiv_trans: "a \ b \ \ a \ b \ a < b" + by (rule not_equiv_le_trans) lemma antisym_conv: "y \ x \ x \ y \ x \ y" by (simp add: equiv_def) end +thm order_trans + +find_theorems "?i < ?j \ ?i \ ?j" + +ML_file \~~/src/Provers/preorder.ML\ + +ML \ +structure Quasi = Quasi_Tac( +struct + +val le_trans = @{thm order_trans}; +val le_refl = @{thm order_refl}; +val eqD1 = @{thm equivD1}; +val eqD2 = @{thm equivD2}; +val less_reflE = @{thm less_irrefl}; +val less_imp_le = @{thm less_imp_le}; +val le_neq_trans = @{thm le_not_equiv_trans}; +val neq_le_trans = @{thm not_equiv_le_trans}; +val less_imp_neq = @{thm less_imp_not_equiv}; + +fun decomp_quasi thy (Const (@{const_name less_eq}, _) $ t1 $ t2) = SOME (t1, "<=", t2) + | decomp_quasi thy (Const (@{const_name less}, _) $ t1 $ t2) = SOME (t1, "<", t2) + | decomp_quasi thy (Const (@{const_name equiv}, _) $ t1 $ t2) = SOME (t1, "=", t2) + | decomp_quasi thy (Const (@{const_name Not}, _) $ (Const (@{const_name equiv}, _) $ t1 $ t2)) = SOME (t1, "~=", t2) + | decomp_quasi thy _ = NONE; + +fun decomp_trans thy t = case decomp_quasi thy t of + x as SOME (t1, "<=", t2) => x + | _ => NONE; + end +); +\ + +end