| author | wenzelm | 
| Tue, 03 Jan 2023 15:32:54 +0100 | |
| changeset 76882 | d9913b41a7bc | 
| parent 74334 | ead56ad40e15 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 63462 | 1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
2  | 
|
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
3  | 
section \<open>A combinator to build partial equivalence relations from a predicate and an equivalence relation\<close>  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
5  | 
theory Combine_PER  | 
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
65366 
diff
changeset
 | 
6  | 
imports Main  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
begin  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
8  | 
|
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
65366 
diff
changeset
 | 
9  | 
unbundle lattice_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
65366 
diff
changeset
 | 
10  | 
|
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
11  | 
definition combine_per :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 63462 | 12  | 
where "combine_per P R = (\<lambda>x y. P x \<and> P y) \<sqinter> R"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
14  | 
lemma combine_per_simp [simp]:  | 
| 63462 | 15  | 
"combine_per P R x y \<longleftrightarrow> P x \<and> P y \<and> x \<approx> y" for R (infixl "\<approx>" 50)  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
16  | 
by (simp add: combine_per_def)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
17  | 
|
| 63462 | 18  | 
lemma combine_per_top [simp]: "combine_per \<top> R = R"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
by (simp add: fun_eq_iff)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
20  | 
|
| 63462 | 21  | 
lemma combine_per_eq [simp]: "combine_per P HOL.eq = HOL.eq \<sqinter> (\<lambda>x y. P x)"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
22  | 
by (auto simp add: fun_eq_iff)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
23  | 
|
| 63462 | 24  | 
lemma symp_combine_per: "symp R \<Longrightarrow> symp (combine_per P R)"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
25  | 
by (auto simp add: symp_def sym_def combine_per_def)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
26  | 
|
| 63462 | 27  | 
lemma transp_combine_per: "transp R \<Longrightarrow> transp (combine_per P R)"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
28  | 
by (auto simp add: transp_def trans_def combine_per_def)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
29  | 
|
| 63462 | 30  | 
lemma combine_perI: "P x \<Longrightarrow> P y \<Longrightarrow> x \<approx> y \<Longrightarrow> combine_per P R x y" for R (infixl "\<approx>" 50)  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
31  | 
by (simp add: combine_per_def)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
32  | 
|
| 63462 | 33  | 
lemma symp_combine_per_symp: "symp R \<Longrightarrow> symp (combine_per P R)"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
34  | 
by (auto intro!: sympI elim: sympE)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
35  | 
|
| 63462 | 36  | 
lemma transp_combine_per_transp: "transp R \<Longrightarrow> transp (combine_per P R)"  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
37  | 
by (auto intro!: transpI elim: transpE)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
38  | 
|
| 63415 | 39  | 
lemma equivp_combine_per_part_equivp [intro?]:  | 
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
40  | 
fixes R (infixl "\<approx>" 50)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
41  | 
assumes "\<exists>x. P x" and "equivp R"  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
42  | 
shows "part_equivp (combine_per P R)"  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
43  | 
proof -  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
44  | 
from \<open>\<exists>x. P x\<close> obtain x where "P x" ..  | 
| 63462 | 45  | 
moreover from \<open>equivp R\<close> have "x \<approx> x"  | 
46  | 
by (rule equivp_reflp)  | 
|
47  | 
ultimately have "\<exists>x. P x \<and> x \<approx> x"  | 
|
48  | 
by blast  | 
|
| 
63377
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
49  | 
with \<open>equivp R\<close> show ?thesis  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
50  | 
by (auto intro!: part_equivpI symp_combine_per_symp transp_combine_per_transp  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
51  | 
elim: equivpE)  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
52  | 
qed  | 
| 
 
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
 
haftmann 
parents:  
diff
changeset
 | 
53  | 
|
| 64267 | 54  | 
end  |