| author | paulson <lp15@cam.ac.uk> | 
| Fri, 28 Feb 2025 13:50:18 +0000 | |
| changeset 82218 | cbf9f856d3e0 | 
| parent 81125 | ec121999a9cb | 
| permissions | -rw-r--r-- | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1  | 
(* Title: HOL/Boolean_Algebras.thy  | 
| 74101 | 2  | 
Author: Brian Huffman  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
3  | 
Author: Florian Haftmann  | 
| 74101 | 4  | 
*)  | 
5  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
6  | 
section \<open>Boolean Algebras\<close>  | 
| 74101 | 7  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
8  | 
theory Boolean_Algebras  | 
| 74101 | 9  | 
imports Lattices  | 
10  | 
begin  | 
|
11  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
12  | 
subsection \<open>Abstract boolean algebra\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
13  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
14  | 
locale abstract_boolean_algebra = conj: abel_semigroup \<open>(\<^bold>\<sqinter>)\<close> + disj: abel_semigroup \<open>(\<^bold>\<squnion>)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
15  | 
for conj :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>\<^bold>\<sqinter>\<close> 70)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
16  | 
and disj :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>\<^bold>\<squnion>\<close> 65) +  | 
| 81125 | 17  | 
fixes compl :: \<open>'a \<Rightarrow> 'a\<close> (\<open>(\<open>open_block notation=\<open>prefix \<^bold>-\<close>\<close>\<^bold>- _)\<close> [81] 80)  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
18  | 
and zero :: \<open>'a\<close> (\<open>\<^bold>0\<close>)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
19  | 
and one :: \<open>'a\<close> (\<open>\<^bold>1\<close>)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
20  | 
assumes conj_disj_distrib: \<open>x \<^bold>\<sqinter> (y \<^bold>\<squnion> z) = (x \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
21  | 
and disj_conj_distrib: \<open>x \<^bold>\<squnion> (y \<^bold>\<sqinter> z) = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (x \<^bold>\<squnion> z)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
22  | 
and conj_one_right: \<open>x \<^bold>\<sqinter> \<^bold>1 = x\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
23  | 
and disj_zero_right: \<open>x \<^bold>\<squnion> \<^bold>0 = x\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
24  | 
and conj_cancel_right [simp]: \<open>x \<^bold>\<sqinter> \<^bold>- x = \<^bold>0\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
25  | 
and disj_cancel_right [simp]: \<open>x \<^bold>\<squnion> \<^bold>- x = \<^bold>1\<close>  | 
| 74101 | 26  | 
begin  | 
27  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
28  | 
sublocale conj: semilattice_neutr \<open>(\<^bold>\<sqinter>)\<close> \<open>\<^bold>1\<close>  | 
| 74101 | 29  | 
proof  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
30  | 
show "x \<^bold>\<sqinter> \<^bold>1 = x" for x  | 
| 74101 | 31  | 
by (fact conj_one_right)  | 
32  | 
show "x \<^bold>\<sqinter> x = x" for x  | 
|
33  | 
proof -  | 
|
34  | 
have "x \<^bold>\<sqinter> x = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> \<^bold>0"  | 
|
35  | 
by (simp add: disj_zero_right)  | 
|
36  | 
also have "\<dots> = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- x)"  | 
|
37  | 
by simp  | 
|
38  | 
also have "\<dots> = x \<^bold>\<sqinter> (x \<^bold>\<squnion> \<^bold>- x)"  | 
|
39  | 
by (simp only: conj_disj_distrib)  | 
|
40  | 
also have "\<dots> = x \<^bold>\<sqinter> \<^bold>1"  | 
|
41  | 
by simp  | 
|
42  | 
also have "\<dots> = x"  | 
|
43  | 
by (simp add: conj_one_right)  | 
|
44  | 
finally show ?thesis .  | 
|
45  | 
qed  | 
|
46  | 
qed  | 
|
47  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
48  | 
sublocale disj: semilattice_neutr \<open>(\<^bold>\<squnion>)\<close> \<open>\<^bold>0\<close>  | 
| 74101 | 49  | 
proof  | 
50  | 
show "x \<^bold>\<squnion> \<^bold>0 = x" for x  | 
|
51  | 
by (fact disj_zero_right)  | 
|
52  | 
show "x \<^bold>\<squnion> x = x" for x  | 
|
53  | 
proof -  | 
|
54  | 
have "x \<^bold>\<squnion> x = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> \<^bold>1"  | 
|
55  | 
by simp  | 
|
56  | 
also have "\<dots> = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> (x \<^bold>\<squnion> \<^bold>- x)"  | 
|
57  | 
by simp  | 
|
58  | 
also have "\<dots> = x \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- x)"  | 
|
59  | 
by (simp only: disj_conj_distrib)  | 
|
60  | 
also have "\<dots> = x \<^bold>\<squnion> \<^bold>0"  | 
|
61  | 
by simp  | 
|
62  | 
also have "\<dots> = x"  | 
|
63  | 
by (simp add: disj_zero_right)  | 
|
64  | 
finally show ?thesis .  | 
|
65  | 
qed  | 
|
66  | 
qed  | 
|
67  | 
||
68  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
69  | 
subsubsection \<open>Complement\<close>  | 
| 74101 | 70  | 
|
71  | 
lemma complement_unique:  | 
|
72  | 
assumes 1: "a \<^bold>\<sqinter> x = \<^bold>0"  | 
|
73  | 
assumes 2: "a \<^bold>\<squnion> x = \<^bold>1"  | 
|
74  | 
assumes 3: "a \<^bold>\<sqinter> y = \<^bold>0"  | 
|
75  | 
assumes 4: "a \<^bold>\<squnion> y = \<^bold>1"  | 
|
76  | 
shows "x = y"  | 
|
77  | 
proof -  | 
|
78  | 
from 1 3 have "(a \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (a \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> y)"  | 
|
79  | 
by simp  | 
|
80  | 
then have "(x \<^bold>\<sqinter> a) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (y \<^bold>\<sqinter> a) \<^bold>\<squnion> (y \<^bold>\<sqinter> x)"  | 
|
81  | 
by (simp add: ac_simps)  | 
|
82  | 
then have "x \<^bold>\<sqinter> (a \<^bold>\<squnion> y) = y \<^bold>\<sqinter> (a \<^bold>\<squnion> x)"  | 
|
83  | 
by (simp add: conj_disj_distrib)  | 
|
84  | 
with 2 4 have "x \<^bold>\<sqinter> \<^bold>1 = y \<^bold>\<sqinter> \<^bold>1"  | 
|
85  | 
by simp  | 
|
86  | 
then show "x = y"  | 
|
87  | 
by simp  | 
|
88  | 
qed  | 
|
89  | 
||
90  | 
lemma compl_unique: "x \<^bold>\<sqinter> y = \<^bold>0 \<Longrightarrow> x \<^bold>\<squnion> y = \<^bold>1 \<Longrightarrow> \<^bold>- x = y"  | 
|
91  | 
by (rule complement_unique [OF conj_cancel_right disj_cancel_right])  | 
|
92  | 
||
93  | 
lemma double_compl [simp]: "\<^bold>- (\<^bold>- x) = x"  | 
|
94  | 
proof (rule compl_unique)  | 
|
95  | 
show "\<^bold>- x \<^bold>\<sqinter> x = \<^bold>0"  | 
|
96  | 
by (simp only: conj_cancel_right conj.commute)  | 
|
97  | 
show "\<^bold>- x \<^bold>\<squnion> x = \<^bold>1"  | 
|
98  | 
by (simp only: disj_cancel_right disj.commute)  | 
|
99  | 
qed  | 
|
100  | 
||
101  | 
lemma compl_eq_compl_iff [simp]:  | 
|
102  | 
\<open>\<^bold>- x = \<^bold>- y \<longleftrightarrow> x = y\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)  | 
|
103  | 
proof  | 
|
104  | 
assume \<open>?Q\<close>  | 
|
105  | 
then show ?P by simp  | 
|
106  | 
next  | 
|
107  | 
assume \<open>?P\<close>  | 
|
108  | 
then have \<open>\<^bold>- (\<^bold>- x) = \<^bold>- (\<^bold>- y)\<close>  | 
|
109  | 
by simp  | 
|
110  | 
then show ?Q  | 
|
111  | 
by simp  | 
|
112  | 
qed  | 
|
113  | 
||
114  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
115  | 
subsubsection \<open>Conjunction\<close>  | 
| 74101 | 116  | 
|
117  | 
lemma conj_zero_right [simp]: "x \<^bold>\<sqinter> \<^bold>0 = \<^bold>0"  | 
|
118  | 
using conj.left_idem conj_cancel_right by fastforce  | 
|
119  | 
||
120  | 
lemma compl_one [simp]: "\<^bold>- \<^bold>1 = \<^bold>0"  | 
|
121  | 
by (rule compl_unique [OF conj_zero_right disj_zero_right])  | 
|
122  | 
||
123  | 
lemma conj_zero_left [simp]: "\<^bold>0 \<^bold>\<sqinter> x = \<^bold>0"  | 
|
124  | 
by (subst conj.commute) (rule conj_zero_right)  | 
|
125  | 
||
126  | 
lemma conj_cancel_left [simp]: "\<^bold>- x \<^bold>\<sqinter> x = \<^bold>0"  | 
|
127  | 
by (subst conj.commute) (rule conj_cancel_right)  | 
|
128  | 
||
129  | 
lemma conj_disj_distrib2: "(y \<^bold>\<squnion> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x)"  | 
|
130  | 
by (simp only: conj.commute conj_disj_distrib)  | 
|
131  | 
||
132  | 
lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2  | 
|
133  | 
||
134  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
135  | 
subsubsection \<open>Disjunction\<close>  | 
| 74101 | 136  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
137  | 
context  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
138  | 
begin  | 
| 74101 | 139  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
140  | 
interpretation dual: abstract_boolean_algebra \<open>(\<^bold>\<squnion>)\<close> \<open>(\<^bold>\<sqinter>)\<close> compl \<open>\<^bold>1\<close> \<open>\<^bold>0\<close>  | 
| 74101 | 141  | 
apply standard  | 
142  | 
apply (rule disj_conj_distrib)  | 
|
143  | 
apply (rule conj_disj_distrib)  | 
|
144  | 
apply simp_all  | 
|
145  | 
done  | 
|
146  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
147  | 
lemma disj_one_right [simp]: "x \<^bold>\<squnion> \<^bold>1 = \<^bold>1"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
148  | 
by (fact dual.conj_zero_right)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
149  | 
|
| 74101 | 150  | 
lemma compl_zero [simp]: "\<^bold>- \<^bold>0 = \<^bold>1"  | 
151  | 
by (fact dual.compl_one)  | 
|
152  | 
||
153  | 
lemma disj_one_left [simp]: "\<^bold>1 \<^bold>\<squnion> x = \<^bold>1"  | 
|
154  | 
by (fact dual.conj_zero_left)  | 
|
155  | 
||
156  | 
lemma disj_cancel_left [simp]: "\<^bold>- x \<^bold>\<squnion> x = \<^bold>1"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
157  | 
by (fact dual.conj_cancel_left)  | 
| 74101 | 158  | 
|
159  | 
lemma disj_conj_distrib2: "(y \<^bold>\<sqinter> z) \<^bold>\<squnion> x = (y \<^bold>\<squnion> x) \<^bold>\<sqinter> (z \<^bold>\<squnion> x)"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
160  | 
by (fact dual.conj_disj_distrib2)  | 
| 74101 | 161  | 
|
162  | 
lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2  | 
|
163  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
164  | 
end  | 
| 74101 | 165  | 
|
166  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
167  | 
subsubsection \<open>De Morgan's Laws\<close>  | 
| 74101 | 168  | 
|
169  | 
lemma de_Morgan_conj [simp]: "\<^bold>- (x \<^bold>\<sqinter> y) = \<^bold>- x \<^bold>\<squnion> \<^bold>- y"  | 
|
170  | 
proof (rule compl_unique)  | 
|
171  | 
have "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<^bold>- y)"  | 
|
172  | 
by (rule conj_disj_distrib)  | 
|
173  | 
also have "\<dots> = (y \<^bold>\<sqinter> (x \<^bold>\<sqinter> \<^bold>- x)) \<^bold>\<squnion> (x \<^bold>\<sqinter> (y \<^bold>\<sqinter> \<^bold>- y))"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
174  | 
by (simp only: ac_simps)  | 
| 74101 | 175  | 
finally show "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = \<^bold>0"  | 
176  | 
by (simp only: conj_cancel_right conj_zero_right disj_zero_right)  | 
|
177  | 
next  | 
|
178  | 
have "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = (x \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)) \<^bold>\<sqinter> (y \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y))"  | 
|
179  | 
by (rule disj_conj_distrib2)  | 
|
180  | 
also have "\<dots> = (\<^bold>- y \<^bold>\<squnion> (x \<^bold>\<squnion> \<^bold>- x)) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> (y \<^bold>\<squnion> \<^bold>- y))"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
181  | 
by (simp only: ac_simps)  | 
| 74101 | 182  | 
finally show "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = \<^bold>1"  | 
183  | 
by (simp only: disj_cancel_right disj_one_right conj_one_right)  | 
|
184  | 
qed  | 
|
185  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
186  | 
context  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
187  | 
begin  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
188  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
189  | 
interpretation dual: abstract_boolean_algebra \<open>(\<^bold>\<squnion>)\<close> \<open>(\<^bold>\<sqinter>)\<close> compl \<open>\<^bold>1\<close> \<open>\<^bold>0\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
190  | 
apply standard  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
191  | 
apply (rule disj_conj_distrib)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
192  | 
apply (rule conj_disj_distrib)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
193  | 
apply simp_all  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
194  | 
done  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
195  | 
|
| 74101 | 196  | 
lemma de_Morgan_disj [simp]: "\<^bold>- (x \<^bold>\<squnion> y) = \<^bold>- x \<^bold>\<sqinter> \<^bold>- y"  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
197  | 
by (fact dual.de_Morgan_conj)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
198  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
199  | 
end  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
200  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
201  | 
end  | 
| 74101 | 202  | 
|
203  | 
||
204  | 
subsection \<open>Symmetric Difference\<close>  | 
|
205  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
206  | 
locale abstract_boolean_algebra_sym_diff = abstract_boolean_algebra +  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
207  | 
fixes xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>\<^bold>\<ominus>\<close> 65)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
208  | 
assumes xor_def : \<open>x \<^bold>\<ominus> y = (x \<^bold>\<sqinter> \<^bold>- y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> y)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
209  | 
begin  | 
| 74101 | 210  | 
|
211  | 
sublocale xor: comm_monoid xor \<open>\<^bold>0\<close>  | 
|
212  | 
proof  | 
|
213  | 
fix x y z :: 'a  | 
|
214  | 
let ?t = "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z)"  | 
|
215  | 
have "?t \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (z \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- y) = ?t \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z \<^bold>\<sqinter> \<^bold>- z)"  | 
|
216  | 
by (simp only: conj_cancel_right conj_zero_right)  | 
|
217  | 
then show "(x \<^bold>\<ominus> y) \<^bold>\<ominus> z = x \<^bold>\<ominus> (y \<^bold>\<ominus> z)"  | 
|
218  | 
by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
219  | 
(simp only: conj_disj_distribs conj_ac ac_simps)  | 
| 74101 | 220  | 
show "x \<^bold>\<ominus> y = y \<^bold>\<ominus> x"  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
221  | 
by (simp only: xor_def ac_simps)  | 
| 74101 | 222  | 
show "x \<^bold>\<ominus> \<^bold>0 = x"  | 
223  | 
by (simp add: xor_def)  | 
|
224  | 
qed  | 
|
225  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
226  | 
lemma xor_def2:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
227  | 
\<open>x \<^bold>\<ominus> y = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
228  | 
proof -  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
229  | 
note xor_def [of x y]  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
230  | 
also have \<open>x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<squnion> \<^bold>- x \<^bold>\<sqinter> y = ((x \<^bold>\<squnion> \<^bold>- x) \<^bold>\<sqinter> (\<^bold>- y \<^bold>\<squnion> \<^bold>- x)) \<^bold>\<sqinter> (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<^bold>- y \<^bold>\<squnion> y)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
231  | 
by (simp add: ac_simps disj_conj_distribs)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
232  | 
also have \<open>\<dots> = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
233  | 
by (simp add: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
234  | 
finally show ?thesis .  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
235  | 
qed  | 
| 74101 | 236  | 
|
237  | 
lemma xor_one_right [simp]: "x \<^bold>\<ominus> \<^bold>1 = \<^bold>- x"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
238  | 
by (simp only: xor_def compl_one conj_zero_right conj_one_right disj.left_neutral)  | 
| 74101 | 239  | 
|
240  | 
lemma xor_one_left [simp]: "\<^bold>1 \<^bold>\<ominus> x = \<^bold>- x"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
241  | 
using xor_one_right [of x] by (simp add: ac_simps)  | 
| 74101 | 242  | 
|
243  | 
lemma xor_self [simp]: "x \<^bold>\<ominus> x = \<^bold>0"  | 
|
244  | 
by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)  | 
|
245  | 
||
246  | 
lemma xor_left_self [simp]: "x \<^bold>\<ominus> (x \<^bold>\<ominus> y) = y"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
247  | 
by (simp only: xor.assoc [symmetric] xor_self xor.left_neutral)  | 
| 74101 | 248  | 
|
249  | 
lemma xor_compl_left [simp]: "\<^bold>- x \<^bold>\<ominus> y = \<^bold>- (x \<^bold>\<ominus> y)"  | 
|
250  | 
by (simp add: ac_simps flip: xor_one_left)  | 
|
251  | 
||
252  | 
lemma xor_compl_right [simp]: "x \<^bold>\<ominus> \<^bold>- y = \<^bold>- (x \<^bold>\<ominus> y)"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
253  | 
using xor.commute xor_compl_left by auto  | 
| 74101 | 254  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
255  | 
lemma xor_cancel_right [simp]: "x \<^bold>\<ominus> \<^bold>- x = \<^bold>1"  | 
| 74101 | 256  | 
by (simp only: xor_compl_right xor_self compl_zero)  | 
257  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
258  | 
lemma xor_cancel_left [simp]: "\<^bold>- x \<^bold>\<ominus> x = \<^bold>1"  | 
| 74101 | 259  | 
by (simp only: xor_compl_left xor_self compl_zero)  | 
260  | 
||
261  | 
lemma conj_xor_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<ominus> z) = (x \<^bold>\<sqinter> y) \<^bold>\<ominus> (x \<^bold>\<sqinter> z)"  | 
|
262  | 
proof -  | 
|
263  | 
have *: "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z) =  | 
|
264  | 
(y \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z)"  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
265  | 
by (simp only: conj_cancel_right conj_zero_right disj.left_neutral)  | 
| 74101 | 266  | 
then show "x \<^bold>\<sqinter> (y \<^bold>\<ominus> z) = (x \<^bold>\<sqinter> y) \<^bold>\<ominus> (x \<^bold>\<sqinter> z)"  | 
267  | 
by (simp (no_asm_use) only:  | 
|
268  | 
xor_def de_Morgan_disj de_Morgan_conj double_compl  | 
|
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
269  | 
conj_disj_distribs ac_simps)  | 
| 74101 | 270  | 
qed  | 
271  | 
||
272  | 
lemma conj_xor_distrib2: "(y \<^bold>\<ominus> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<ominus> (z \<^bold>\<sqinter> x)"  | 
|
273  | 
by (simp add: conj.commute conj_xor_distrib)  | 
|
274  | 
||
275  | 
lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2  | 
|
276  | 
||
277  | 
end  | 
|
278  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
279  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
280  | 
subsection \<open>Type classes\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
281  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
282  | 
class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus +  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
283  | 
assumes inf_compl_bot: \<open>x \<sqinter> - x = \<bottom>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
284  | 
and sup_compl_top: \<open>x \<squnion> - x = \<top>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
285  | 
assumes diff_eq: \<open>x - y = x \<sqinter> - y\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
286  | 
begin  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
287  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
288  | 
sublocale boolean_algebra: abstract_boolean_algebra \<open>(\<sqinter>)\<close> \<open>(\<squnion>)\<close> uminus \<bottom> \<top>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
289  | 
apply standard  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
290  | 
apply (rule inf_sup_distrib1)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
291  | 
apply (rule sup_inf_distrib1)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
292  | 
apply (simp_all add: ac_simps inf_compl_bot sup_compl_top)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
293  | 
done  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
294  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
295  | 
lemma compl_inf_bot: "- x \<sqinter> x = \<bottom>"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
296  | 
by (fact boolean_algebra.conj_cancel_left)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
297  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
298  | 
lemma compl_sup_top: "- x \<squnion> x = \<top>"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
299  | 
by (fact boolean_algebra.disj_cancel_left)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
300  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
301  | 
lemma compl_unique:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
302  | 
assumes "x \<sqinter> y = \<bottom>"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
303  | 
and "x \<squnion> y = \<top>"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
304  | 
shows "- x = y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
305  | 
using assms by (rule boolean_algebra.compl_unique)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
306  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
307  | 
lemma double_compl: "- (- x) = x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
308  | 
by (fact boolean_algebra.double_compl)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
309  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
310  | 
lemma compl_eq_compl_iff: "- x = - y \<longleftrightarrow> x = y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
311  | 
by (fact boolean_algebra.compl_eq_compl_iff)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
312  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
313  | 
lemma compl_bot_eq: "- \<bottom> = \<top>"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
314  | 
by (fact boolean_algebra.compl_zero)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
315  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
316  | 
lemma compl_top_eq: "- \<top> = \<bottom>"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
317  | 
by (fact boolean_algebra.compl_one)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
318  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
319  | 
lemma compl_inf: "- (x \<sqinter> y) = - x \<squnion> - y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
320  | 
by (fact boolean_algebra.de_Morgan_conj)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
321  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
322  | 
lemma compl_sup: "- (x \<squnion> y) = - x \<sqinter> - y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
323  | 
by (fact boolean_algebra.de_Morgan_disj)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
324  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
325  | 
lemma compl_mono:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
326  | 
assumes "x \<le> y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
327  | 
shows "- y \<le> - x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
328  | 
proof -  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
329  | 
from assms have "x \<squnion> y = y" by (simp only: le_iff_sup)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
330  | 
then have "- (x \<squnion> y) = - y" by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
331  | 
then have "- x \<sqinter> - y = - y" by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
332  | 
then have "- y \<sqinter> - x = - y" by (simp only: inf_commute)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
333  | 
then show ?thesis by (simp only: le_iff_inf)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
334  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
335  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
336  | 
lemma compl_le_compl_iff [simp]: "- x \<le> - y \<longleftrightarrow> y \<le> x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
337  | 
by (auto dest: compl_mono)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
338  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
339  | 
lemma compl_le_swap1:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
340  | 
assumes "y \<le> - x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
341  | 
shows "x \<le> -y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
342  | 
proof -  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
343  | 
from assms have "- (- x) \<le> - y" by (simp only: compl_le_compl_iff)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
344  | 
then show ?thesis by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
345  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
346  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
347  | 
lemma compl_le_swap2:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
348  | 
assumes "- y \<le> x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
349  | 
shows "- x \<le> y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
350  | 
proof -  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
351  | 
from assms have "- x \<le> - (- y)" by (simp only: compl_le_compl_iff)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
352  | 
then show ?thesis by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
353  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
354  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
355  | 
lemma compl_less_compl_iff [simp]: "- x < - y \<longleftrightarrow> y < x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
356  | 
by (auto simp add: less_le)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
357  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
358  | 
lemma compl_less_swap1:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
359  | 
assumes "y < - x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
360  | 
shows "x < - y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
361  | 
proof -  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
362  | 
from assms have "- (- x) < - y" by (simp only: compl_less_compl_iff)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
363  | 
then show ?thesis by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
364  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
365  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
366  | 
lemma compl_less_swap2:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
367  | 
assumes "- y < x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
368  | 
shows "- x < y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
369  | 
proof -  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
370  | 
from assms have "- x < - (- y)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
371  | 
by (simp only: compl_less_compl_iff)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
372  | 
then show ?thesis by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
373  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
374  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
375  | 
lemma sup_cancel_left1: \<open>x \<squnion> a \<squnion> (- x \<squnion> b) = \<top>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
376  | 
by (simp add: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
377  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
378  | 
lemma sup_cancel_left2: \<open>- x \<squnion> a \<squnion> (x \<squnion> b) = \<top>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
379  | 
by (simp add: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
380  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
381  | 
lemma inf_cancel_left1: \<open>x \<sqinter> a \<sqinter> (- x \<sqinter> b) = \<bottom>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
382  | 
by (simp add: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
383  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
384  | 
lemma inf_cancel_left2: \<open>- x \<sqinter> a \<sqinter> (x \<sqinter> b) = \<bottom>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
385  | 
by (simp add: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
386  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
387  | 
lemma sup_compl_top_left1 [simp]: \<open>- x \<squnion> (x \<squnion> y) = \<top>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
388  | 
by (simp add: sup_assoc [symmetric])  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
389  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
390  | 
lemma sup_compl_top_left2 [simp]: \<open>x \<squnion> (- x \<squnion> y) = \<top>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
391  | 
using sup_compl_top_left1 [of "- x" y] by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
392  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
393  | 
lemma inf_compl_bot_left1 [simp]: \<open>- x \<sqinter> (x \<sqinter> y) = \<bottom>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
394  | 
by (simp add: inf_assoc [symmetric])  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
395  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
396  | 
lemma inf_compl_bot_left2 [simp]: \<open>x \<sqinter> (- x \<sqinter> y) = \<bottom>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
397  | 
using inf_compl_bot_left1 [of "- x" y] by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
398  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
399  | 
lemma inf_compl_bot_right [simp]: \<open>x \<sqinter> (y \<sqinter> - x) = \<bottom>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
400  | 
by (subst inf_left_commute) simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
401  | 
|
| 74101 | 402  | 
end  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
403  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
404  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
405  | 
subsection \<open>Lattice on \<^typ>\<open>bool\<close>\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
406  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
407  | 
instantiation bool :: boolean_algebra  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
408  | 
begin  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
409  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
410  | 
definition bool_Compl_def [simp]: "uminus = Not"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
411  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
412  | 
definition bool_diff_def [simp]: "A - B \<longleftrightarrow> A \<and> \<not> B"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
413  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
414  | 
definition [simp]: "P \<sqinter> Q \<longleftrightarrow> P \<and> Q"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
415  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
416  | 
definition [simp]: "P \<squnion> Q \<longleftrightarrow> P \<or> Q"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
417  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
418  | 
instance by standard auto  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
419  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
420  | 
end  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
421  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
422  | 
lemma sup_boolI1: "P \<Longrightarrow> P \<squnion> Q"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
423  | 
by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
424  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
425  | 
lemma sup_boolI2: "Q \<Longrightarrow> P \<squnion> Q"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
426  | 
by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
427  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
428  | 
lemma sup_boolE: "P \<squnion> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
429  | 
by auto  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
430  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
431  | 
instance "fun" :: (type, boolean_algebra) boolean_algebra  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
432  | 
by standard (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
433  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
434  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
435  | 
subsection \<open>Lattice on unary and binary predicates\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
436  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
437  | 
lemma inf1I: "A x \<Longrightarrow> B x \<Longrightarrow> (A \<sqinter> B) x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
438  | 
by (simp add: inf_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
439  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
440  | 
lemma inf2I: "A x y \<Longrightarrow> B x y \<Longrightarrow> (A \<sqinter> B) x y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
441  | 
by (simp add: inf_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
442  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
443  | 
lemma inf1E: "(A \<sqinter> B) x \<Longrightarrow> (A x \<Longrightarrow> B x \<Longrightarrow> P) \<Longrightarrow> P"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
444  | 
by (simp add: inf_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
445  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
446  | 
lemma inf2E: "(A \<sqinter> B) x y \<Longrightarrow> (A x y \<Longrightarrow> B x y \<Longrightarrow> P) \<Longrightarrow> P"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
447  | 
by (simp add: inf_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
448  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
449  | 
lemma inf1D1: "(A \<sqinter> B) x \<Longrightarrow> A x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
450  | 
by (rule inf1E)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
451  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
452  | 
lemma inf2D1: "(A \<sqinter> B) x y \<Longrightarrow> A x y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
453  | 
by (rule inf2E)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
454  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
455  | 
lemma inf1D2: "(A \<sqinter> B) x \<Longrightarrow> B x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
456  | 
by (rule inf1E)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
457  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
458  | 
lemma inf2D2: "(A \<sqinter> B) x y \<Longrightarrow> B x y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
459  | 
by (rule inf2E)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
460  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
461  | 
lemma sup1I1: "A x \<Longrightarrow> (A \<squnion> B) x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
462  | 
by (simp add: sup_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
463  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
464  | 
lemma sup2I1: "A x y \<Longrightarrow> (A \<squnion> B) x y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
465  | 
by (simp add: sup_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
466  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
467  | 
lemma sup1I2: "B x \<Longrightarrow> (A \<squnion> B) x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
468  | 
by (simp add: sup_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
469  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
470  | 
lemma sup2I2: "B x y \<Longrightarrow> (A \<squnion> B) x y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
471  | 
by (simp add: sup_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
472  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
473  | 
lemma sup1E: "(A \<squnion> B) x \<Longrightarrow> (A x \<Longrightarrow> P) \<Longrightarrow> (B x \<Longrightarrow> P) \<Longrightarrow> P"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
474  | 
by (simp add: sup_fun_def) iprover  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
475  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
476  | 
lemma sup2E: "(A \<squnion> B) x y \<Longrightarrow> (A x y \<Longrightarrow> P) \<Longrightarrow> (B x y \<Longrightarrow> P) \<Longrightarrow> P"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
477  | 
by (simp add: sup_fun_def) iprover  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
478  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
479  | 
text \<open> \<^medskip> Classical introduction rule: no commitment to \<open>A\<close> vs \<open>B\<close>.\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
480  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
481  | 
lemma sup1CI: "(\<not> B x \<Longrightarrow> A x) \<Longrightarrow> (A \<squnion> B) x"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
482  | 
by (auto simp add: sup_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
483  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
484  | 
lemma sup2CI: "(\<not> B x y \<Longrightarrow> A x y) \<Longrightarrow> (A \<squnion> B) x y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
485  | 
by (auto simp add: sup_fun_def)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
486  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
487  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
488  | 
subsection \<open>Simproc setup\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
489  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
490  | 
locale boolean_algebra_cancel  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
491  | 
begin  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
492  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
493  | 
lemma sup1: "(A::'a::semilattice_sup) \<equiv> sup k a \<Longrightarrow> sup A b \<equiv> sup k (sup a b)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
494  | 
by (simp only: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
495  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
496  | 
lemma sup2: "(B::'a::semilattice_sup) \<equiv> sup k b \<Longrightarrow> sup a B \<equiv> sup k (sup a b)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
497  | 
by (simp only: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
498  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
499  | 
lemma sup0: "(a::'a::bounded_semilattice_sup_bot) \<equiv> sup a bot"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
500  | 
by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
501  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
502  | 
lemma inf1: "(A::'a::semilattice_inf) \<equiv> inf k a \<Longrightarrow> inf A b \<equiv> inf k (inf a b)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
503  | 
by (simp only: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
504  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
505  | 
lemma inf2: "(B::'a::semilattice_inf) \<equiv> inf k b \<Longrightarrow> inf a B \<equiv> inf k (inf a b)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
506  | 
by (simp only: ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
507  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
508  | 
lemma inf0: "(a::'a::bounded_semilattice_inf_top) \<equiv> inf a top"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
509  | 
by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
510  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
511  | 
end  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
512  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
513  | 
ML_file \<open>Tools/boolean_algebra_cancel.ML\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
514  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
515  | 
simproc_setup boolean_algebra_cancel_sup ("sup a b::'a::boolean_algebra") =
 | 
| 
78099
 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 
wenzelm 
parents: 
74123 
diff
changeset
 | 
516  | 
\<open>K (K (try Boolean_Algebra_Cancel.cancel_sup_conv))\<close>  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
517  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
518  | 
simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") =
 | 
| 
78099
 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 
wenzelm 
parents: 
74123 
diff
changeset
 | 
519  | 
\<open>K (K (try Boolean_Algebra_Cancel.cancel_inf_conv))\<close>  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
520  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
521  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
522  | 
context boolean_algebra  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
523  | 
begin  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
524  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
525  | 
lemma shunt1: "(x \<sqinter> y \<le> z) \<longleftrightarrow> (x \<le> -y \<squnion> z)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
526  | 
proof  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
527  | 
assume "x \<sqinter> y \<le> z"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
528  | 
hence "-y \<squnion> (x \<sqinter> y) \<le> -y \<squnion> z"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
529  | 
using sup.mono by blast  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
530  | 
hence "-y \<squnion> x \<le> -y \<squnion> z"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
531  | 
by (simp add: sup_inf_distrib1)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
532  | 
thus "x \<le> -y \<squnion> z"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
533  | 
by simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
534  | 
next  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
535  | 
assume "x \<le> -y \<squnion> z"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
536  | 
hence "x \<sqinter> y \<le> (-y \<squnion> z) \<sqinter> y"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
537  | 
using inf_mono by auto  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
538  | 
thus "x \<sqinter> y \<le> z"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
539  | 
using inf.boundedE inf_sup_distrib2 by auto  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
540  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
541  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
542  | 
lemma shunt2: "(x \<sqinter> -y \<le> z) \<longleftrightarrow> (x \<le> y \<squnion> z)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
543  | 
by (simp add: shunt1)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
544  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
545  | 
lemma inf_shunt: "(x \<sqinter> y = \<bottom>) \<longleftrightarrow> (x \<le> - y)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
546  | 
by (simp add: order.eq_iff shunt1)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
547  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
548  | 
lemma sup_shunt: "(x \<squnion> y = \<top>) \<longleftrightarrow> (- x \<le> y)"  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
549  | 
using inf_shunt [of \<open>- x\<close> \<open>- y\<close>, symmetric]  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
550  | 
by (simp flip: compl_sup compl_top_eq)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
551  | 
|
| 79099 | 552  | 
lemma diff_shunt_var[simp]: "(x - y = \<bottom>) \<longleftrightarrow> (x \<le> y)"  | 
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
553  | 
by (simp add: diff_eq inf_shunt)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
554  | 
|
| 79099 | 555  | 
lemma diff_shunt[simp]: "(\<bottom> = x - y) \<longleftrightarrow> (x \<le> y)"  | 
556  | 
by (auto simp flip: diff_shunt_var)  | 
|
557  | 
||
| 
74123
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
558  | 
lemma sup_neg_inf:  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
559  | 
\<open>p \<le> q \<squnion> r \<longleftrightarrow> p \<sqinter> -q \<le> r\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
560  | 
proof  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
561  | 
assume ?P  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
562  | 
then have \<open>p \<sqinter> - q \<le> (q \<squnion> r) \<sqinter> - q\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
563  | 
by (rule inf_mono) simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
564  | 
then show ?Q  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
565  | 
by (simp add: inf_sup_distrib2)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
566  | 
next  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
567  | 
assume ?Q  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
568  | 
then have \<open>p \<sqinter> - q \<squnion> q \<le> r \<squnion> q\<close>  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
569  | 
by (rule sup_mono) simp  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
570  | 
then show ?P  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
571  | 
by (simp add: sup_inf_distrib ac_simps)  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
572  | 
qed  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
573  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
574  | 
end  | 
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
575  | 
|
| 
 
7c5842b06114
clarified abstract and concrete boolean algebras
 
haftmann 
parents: 
74101 
diff
changeset
 | 
576  | 
end  |