# HG changeset patch # User haftmann # Date 1467654380 -7200 # Node ID 64adf4ba952622335e050062a6342cfb5b823607 # Parent 4c0cc2b356f0b1c04d8fe18c57d97c7819ad5848 combinator to build partial equivalence relations from a predicate and an equivalenc relation diff -r 4c0cc2b356f0 -r 64adf4ba9526 NEWS --- a/NEWS Mon Jul 04 19:46:20 2016 +0200 +++ b/NEWS Mon Jul 04 19:46:20 2016 +0200 @@ -136,6 +136,9 @@ *** HOL *** +* Theory Library/Combinator_PER.thy: combinator to build partial +equivalence relations from a predicate and an equivalenc relation. + * Theory Library/Perm.thy: basic facts about almost everywhere fix bijections. diff -r 4c0cc2b356f0 -r 64adf4ba9526 src/HOL/Library/Combine_PER.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Combine_PER.thy Mon Jul 04 19:46:20 2016 +0200 @@ -0,0 +1,60 @@ +(* Author: Florian Haftmann, TU Muenchen *) + +section \A combinator to build partial equivalence relations from a predicate and an equivalence relation\ + +theory Combine_PER +imports Main "~~/src/HOL/Library/Lattice_Syntax" +begin + +definition combine_per :: "('a \ bool) \ ('a \ 'a \ bool) \ 'a \ 'a \ bool" +where + "combine_per P R = (\x y. P x \ P y) \ R" + +lemma combine_per_simp [simp]: + fixes R (infixl "\" 50) + shows "combine_per P R x y \ P x \ P y \ x \ y" + by (simp add: combine_per_def) + +lemma combine_per_top [simp]: + "combine_per \ R = R" + by (simp add: fun_eq_iff) + +lemma combine_per_eq [simp]: + "combine_per P HOL.eq = HOL.eq \ (\x y. P x)" + by (auto simp add: fun_eq_iff) + +lemma symp_combine_per: + "symp R \ symp (combine_per P R)" + by (auto simp add: symp_def sym_def combine_per_def) + +lemma transp_combine_per: + "transp R \ transp (combine_per P R)" + by (auto simp add: transp_def trans_def combine_per_def) + +lemma combine_perI: + fixes R (infixl "\" 50) + shows "P x \ P y \ x \ y \ combine_per P R x y" + by (simp add: combine_per_def) + +lemma symp_combine_per_symp: + "symp R \ symp (combine_per P R)" + by (auto intro!: sympI elim: sympE) + +lemma transp_combine_per_transp: + "transp R \ transp (combine_per P R)" + by (auto intro!: transpI elim: transpE) + +lemma equivp_combine_per_part_equivp: + fixes R (infixl "\" 50) + assumes "\x. P x" and "equivp R" + shows "part_equivp (combine_per P R)" +proof - + from \\x. P x\ obtain x where "P x" .. + moreover from \equivp R\ have "x \ x" by (rule equivp_reflp) + ultimately have "\x. P x \ x \ x" by blast + with \equivp R\ show ?thesis + by (auto intro!: part_equivpI symp_combine_per_symp transp_combine_per_transp + elim: equivpE) +qed + +end \ No newline at end of file diff -r 4c0cc2b356f0 -r 64adf4ba9526 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Mon Jul 04 19:46:20 2016 +0200 +++ b/src/HOL/Library/Library.thy Mon Jul 04 19:46:20 2016 +0200 @@ -12,6 +12,7 @@ Code_Test ContNotDenum Convex + Combine_PER Complete_Partial_Order2 Countable Countable_Complete_Lattices