# HG changeset patch # User desharna # Date 1669036290 -3600 # Node ID 4352d0ff165a3b6156f186417ddd6e624b1da3e1 # Parent a7d9e34c85e6bdc9addeacf734c1b2f6e811056c introduced predicates irrefl_on and irreflp_on, and redefined irrefl and irreflp as abbreviations diff -r a7d9e34c85e6 -r 4352d0ff165a NEWS --- a/NEWS Mon Nov 21 18:24:55 2022 +0100 +++ b/NEWS Mon Nov 21 14:11:30 2022 +0100 @@ -26,14 +26,21 @@ Minor INCOMPATIBILITY. * Theory "HOL.Relation": + - Added predicates irrefl_on and irreflp_on and redefined irrefl and + irreflp to be abbreviations. Lemmas irrefl_def and irreflp_def are + explicitly provided for backward compatibility but their usage is + discouraged. Minor INCOMPATIBILITY. - Strengthened (and renamed) lemmas. Minor INCOMPATIBILITY. total_on_singleton reflp_equality[simp] ~> reflp_on_equality[simp] - Added lemmas. antisym_if_asymp antisymp_if_asymp - irreflD - irreflpD + irrefl_onD + irrefl_onI + irreflp_onD + irreflp_onI + irreflp_on_irrefl_on_eq[pred_set_conv] linorder.totalp_on_ge[simp] linorder.totalp_on_greater[simp] linorder.totalp_on_le[simp] diff -r a7d9e34c85e6 -r 4352d0ff165a src/HOL/Relation.thy --- a/src/HOL/Relation.thy Mon Nov 21 18:24:55 2022 +0100 +++ b/src/HOL/Relation.thy Mon Nov 21 14:11:30 2022 +0100 @@ -1,6 +1,7 @@ (* Title: HOL/Relation.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer, TU Muenchen + Author: Martin Desharnais, MPI-INF Saarbruecken *) section \Relations -- as sets of pairs, and binary predicates\ @@ -265,29 +266,59 @@ subsubsection \Irreflexivity\ -definition irrefl :: "'a rel \ bool" - where "irrefl r \ (\a. (a, a) \ r)" +definition irrefl_on :: "'a set \ 'a rel \ bool" where + "irrefl_on A r \ (\a \ A. (a, a) \ r)" + +abbreviation irrefl :: "'a rel \ bool" where + "irrefl \ irrefl_on UNIV" + +definition irreflp_on :: "'a set \ ('a \ 'a \ bool) \ bool" where + "irreflp_on A R \ (\a \ A. \ R a a)" -definition irreflp :: "('a \ 'a \ bool) \ bool" - where "irreflp R \ (\a. \ R a a)" +abbreviation irreflp :: "('a \ 'a \ bool) \ bool" where + "irreflp \ irreflp_on UNIV" + +lemma irrefl_def[no_atp]: "irrefl r \ (\a. (a, a) \ r)" + by (simp add: irrefl_on_def) + +lemma irreflp_def[no_atp]: "irreflp R \ (\a. \ R a a)" + by (simp add: irreflp_on_def) + +text \@{thm [source] irrefl_def} and @{thm [source] irreflp_def} are for backward compatibility.\ -lemma irreflp_irrefl_eq [pred_set_conv]: "irreflp (\a b. (a, b) \ R) \ irrefl R" - by (simp add: irrefl_def irreflp_def) +lemma irreflp_on_irrefl_on_eq [pred_set_conv]: "irreflp_on A (\a b. (a, b) \ r) \ irrefl_on A r" + by (simp add: irrefl_on_def irreflp_on_def) + +lemmas irreflp_irrefl_eq = irreflp_on_irrefl_on_eq[of UNIV] + +lemma irrefl_onI: "(\a. a \ A \ (a, a) \ r) \ irrefl_on A r" + by (simp add: irrefl_on_def) -lemma irreflI [intro?]: "(\a. (a, a) \ R) \ irrefl R" - by (simp add: irrefl_def) +lemma irreflI[intro?]: "(\a. (a, a) \ r) \ irrefl r" + by (rule irrefl_onI[of UNIV, simplified]) + +lemma irreflp_onI: "(\a. a \ A \ \ R a a) \ irreflp_on A R" + by (simp add: irreflp_on_def) -lemma irreflpI [intro?]: "(\a. \ R a a) \ irreflp R" - by (fact irreflI [to_pred]) +lemma irreflpI[intro?]: "(\a. \ R a a) \ irreflp R" + by (rule irreflp_onI[of UNIV, simplified]) + +lemma irrefl_onD: "irrefl_on A r \ a \ A \ (a, a) \ r" + by (simp add: irrefl_on_def) lemma irreflD: "irrefl r \ (x, x) \ r" - unfolding irrefl_def by simp + by (rule irrefl_onD[of UNIV, simplified]) + +lemma irreflp_onD: "irreflp_on A R \ a \ A \ \ R a a" + by (simp add: irreflp_on_def) lemma irreflpD: "irreflp R \ \ R x x" - unfolding irreflp_def by simp + by (rule irreflp_onD[of UNIV, simplified]) -lemma irrefl_distinct [code]: "irrefl r \ (\(a, b) \ r. a \ b)" - by (auto simp add: irrefl_def) +lemma irrefl_on_distinct [code]: "irrefl_on A r \ (\(a, b) \ r. a \ A \ b \ A \ a \ b)" + by (auto simp add: irrefl_on_def) + +lemmas irrefl_distinct = irrefl_on_distinct \ \For backward compatibility\ lemma (in preorder) irreflp_less[simp]: "irreflp (<)" by (simp add: irreflpI) diff -r a7d9e34c85e6 -r 4352d0ff165a src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Mon Nov 21 18:24:55 2022 +0100 +++ b/src/HOL/Wellfounded.thy Mon Nov 21 14:11:30 2022 +0100 @@ -75,7 +75,7 @@ using wf_irrefl [OF assms] by (auto simp add: irrefl_def) lemma wfP_imp_irreflp: "wfP r \ irreflp r" - by (rule wf_imp_irrefl[to_pred]) + by (rule wf_imp_irrefl[to_pred, folded top_set_def]) lemma wf_wellorderI: assumes wf: "wf {(x::'a::ord, y). x < y}"