author | wenzelm |
Sun, 29 Sep 2024 21:57:47 +0200 | |
changeset 81011 | 6d34c2bedaa3 |
parent 80923 | 6c9628a116cc |
child 81091 | c007e6d9941d |
permissions | -rw-r--r-- |
1268 | 1 |
(* Title: FOL/IFOL.thy |
11677 | 2 |
Author: Lawrence C Paulson and Markus Wenzel |
3 |
*) |
|
35 | 4 |
|
60770 | 5 |
section \<open>Intuitionistic first-order logic\<close> |
35 | 6 |
|
15481 | 7 |
theory IFOL |
71839 | 8 |
imports Pure |
73015 | 9 |
abbrevs "?<" = "\<exists>\<^sub>\<le>\<^sub>1" |
15481 | 10 |
begin |
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
11 |
|
69605 | 12 |
ML_file \<open>~~/src/Tools/misc_legacy.ML\<close> |
13 |
ML_file \<open>~~/src/Provers/splitter.ML\<close> |
|
14 |
ML_file \<open>~~/src/Provers/hypsubst.ML\<close> |
|
15 |
ML_file \<open>~~/src/Tools/IsaPlanner/zipper.ML\<close> |
|
16 |
ML_file \<open>~~/src/Tools/IsaPlanner/isand.ML\<close> |
|
17 |
ML_file \<open>~~/src/Tools/IsaPlanner/rw_inst.ML\<close> |
|
18 |
ML_file \<open>~~/src/Provers/quantifier1.ML\<close> |
|
19 |
ML_file \<open>~~/src/Tools/intuitionistic.ML\<close> |
|
20 |
ML_file \<open>~~/src/Tools/project_rule.ML\<close> |
|
21 |
ML_file \<open>~~/src/Tools/atomize_elim.ML\<close> |
|
48891 | 22 |
|
0 | 23 |
|
60770 | 24 |
subsection \<open>Syntax and axiomatic basis\<close> |
11677 | 25 |
|
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39159
diff
changeset
|
26 |
setup Pure_Thy.old_appl_syntax_setup |
70880 | 27 |
setup \<open>Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\<close> |
26956
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents:
26580
diff
changeset
|
28 |
|
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
52241
diff
changeset
|
29 |
class "term" |
69590 | 30 |
default_sort \<open>term\<close> |
0 | 31 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
32 |
typedecl o |
79 | 33 |
|
11747 | 34 |
judgment |
80923 | 35 |
Trueprop :: \<open>o \<Rightarrow> prop\<close> (\<open>(\<open>notation=judgment\<close>_)\<close> 5) |
0 | 36 |
|
79 | 37 |
|
60770 | 38 |
subsubsection \<open>Equality\<close> |
35 | 39 |
|
46972 | 40 |
axiomatization |
69590 | 41 |
eq :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>=\<close> 50) |
46972 | 42 |
where |
69590 | 43 |
refl: \<open>a = a\<close> and |
44 |
subst: \<open>a = b \<Longrightarrow> P(a) \<Longrightarrow> P(b)\<close> |
|
79 | 45 |
|
0 | 46 |
|
60770 | 47 |
subsubsection \<open>Propositional logic\<close> |
46972 | 48 |
|
49 |
axiomatization |
|
69590 | 50 |
False :: \<open>o\<close> and |
51 |
conj :: \<open>[o, o] => o\<close> (infixr \<open>\<and>\<close> 35) and |
|
52 |
disj :: \<open>[o, o] => o\<close> (infixr \<open>\<or>\<close> 30) and |
|
53 |
imp :: \<open>[o, o] => o\<close> (infixr \<open>\<longrightarrow>\<close> 25) |
|
46972 | 54 |
where |
69590 | 55 |
conjI: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q\<close> and |
56 |
conjunct1: \<open>P \<and> Q \<Longrightarrow> P\<close> and |
|
57 |
conjunct2: \<open>P \<and> Q \<Longrightarrow> Q\<close> and |
|
46972 | 58 |
|
69590 | 59 |
disjI1: \<open>P \<Longrightarrow> P \<or> Q\<close> and |
60 |
disjI2: \<open>Q \<Longrightarrow> P \<or> Q\<close> and |
|
61 |
disjE: \<open>\<lbrakk>P \<or> Q; P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close> and |
|
46972 | 62 |
|
69590 | 63 |
impI: \<open>(P \<Longrightarrow> Q) \<Longrightarrow> P \<longrightarrow> Q\<close> and |
64 |
mp: \<open>\<lbrakk>P \<longrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close> and |
|
46972 | 65 |
|
69590 | 66 |
FalseE: \<open>False \<Longrightarrow> P\<close> |
46972 | 67 |
|
68 |
||
60770 | 69 |
subsubsection \<open>Quantifiers\<close> |
46972 | 70 |
|
71 |
axiomatization |
|
69590 | 72 |
All :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close> (binder \<open>\<forall>\<close> 10) and |
73 |
Ex :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close> (binder \<open>\<exists>\<close> 10) |
|
46972 | 74 |
where |
69590 | 75 |
allI: \<open>(\<And>x. P(x)) \<Longrightarrow> (\<forall>x. P(x))\<close> and |
76 |
spec: \<open>(\<forall>x. P(x)) \<Longrightarrow> P(x)\<close> and |
|
77 |
exI: \<open>P(x) \<Longrightarrow> (\<exists>x. P(x))\<close> and |
|
78 |
exE: \<open>\<lbrakk>\<exists>x. P(x); \<And>x. P(x) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close> |
|
46972 | 79 |
|
80 |
||
60770 | 81 |
subsubsection \<open>Definitions\<close> |
46972 | 82 |
|
69590 | 83 |
definition \<open>True \<equiv> False \<longrightarrow> False\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
84 |
|
69587 | 85 |
definition Not (\<open>\<not> _\<close> [40] 40) |
69590 | 86 |
where not_def: \<open>\<not> P \<equiv> P \<longrightarrow> False\<close> |
46972 | 87 |
|
69587 | 88 |
definition iff (infixr \<open>\<longleftrightarrow>\<close> 25) |
69590 | 89 |
where \<open>P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
90 |
|
71839 | 91 |
definition Uniq :: "('a \<Rightarrow> o) \<Rightarrow> o" |
92 |
where \<open>Uniq(P) \<equiv> (\<forall>x y. P(x) \<longrightarrow> P(y) \<longrightarrow> y = x)\<close> |
|
93 |
||
69590 | 94 |
definition Ex1 :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close> (binder \<open>\<exists>!\<close> 10) |
95 |
where ex1_def: \<open>\<exists>!x. P(x) \<equiv> \<exists>x. P(x) \<and> (\<forall>y. P(y) \<longrightarrow> y = x)\<close> |
|
46972 | 96 |
|
62020 | 97 |
axiomatization where \<comment> \<open>Reflection, admissible\<close> |
69590 | 98 |
eq_reflection: \<open>(x = y) \<Longrightarrow> (x \<equiv> y)\<close> and |
99 |
iff_reflection: \<open>(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)\<close> |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
100 |
|
69590 | 101 |
abbreviation not_equal :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>\<noteq>\<close> 50) |
102 |
where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close> |
|
46972 | 103 |
|
80917 | 104 |
syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o" (\<open>(\<open>indent=2 notation=\<open>binder \<exists>\<^sub>\<le>\<^sub>1\<close>\<close>\<exists>\<^sub>\<le>\<^sub>1 _./ _)\<close> [0, 10] 10) |
80761 | 105 |
syntax_consts "_Uniq" \<rightleftharpoons> Uniq |
71839 | 106 |
translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)" |
107 |
||
108 |
print_translation \<open> |
|
109 |
[Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\<open>Uniq\<close> \<^syntax_const>\<open>_Uniq\<close>] |
|
110 |
\<close> \<comment> \<open>to avoid eta-contraction of body\<close> |
|
111 |
||
46972 | 112 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
113 |
subsubsection \<open>Old-style ASCII syntax\<close> |
79 | 114 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
115 |
notation (ASCII) |
69587 | 116 |
not_equal (infixl \<open>~=\<close> 50) and |
117 |
Not (\<open>~ _\<close> [40] 40) and |
|
118 |
conj (infixr \<open>&\<close> 35) and |
|
119 |
disj (infixr \<open>|\<close> 30) and |
|
120 |
All (binder \<open>ALL \<close> 10) and |
|
121 |
Ex (binder \<open>EX \<close> 10) and |
|
122 |
Ex1 (binder \<open>EX! \<close> 10) and |
|
123 |
imp (infixr \<open>-->\<close> 25) and |
|
124 |
iff (infixr \<open><->\<close> 25) |
|
35 | 125 |
|
13779 | 126 |
|
60770 | 127 |
subsection \<open>Lemmas and proof tools\<close> |
11677 | 128 |
|
46972 | 129 |
lemmas strip = impI allI |
130 |
||
69590 | 131 |
lemma TrueI: \<open>True\<close> |
21539 | 132 |
unfolding True_def by (rule impI) |
133 |
||
134 |
||
62020 | 135 |
subsubsection \<open>Sequent-style elimination rules for \<open>\<and>\<close> \<open>\<longrightarrow>\<close> and \<open>\<forall>\<close>\<close> |
21539 | 136 |
|
137 |
lemma conjE: |
|
69590 | 138 |
assumes major: \<open>P \<and> Q\<close> |
139 |
and r: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R\<close> |
|
140 |
shows \<open>R\<close> |
|
71839 | 141 |
proof (rule r) |
142 |
show "P" |
|
143 |
by (rule major [THEN conjunct1]) |
|
144 |
show "Q" |
|
145 |
by (rule major [THEN conjunct2]) |
|
146 |
qed |
|
21539 | 147 |
|
148 |
lemma impE: |
|
69590 | 149 |
assumes major: \<open>P \<longrightarrow> Q\<close> |
150 |
and \<open>P\<close> |
|
151 |
and r: \<open>Q \<Longrightarrow> R\<close> |
|
152 |
shows \<open>R\<close> |
|
71839 | 153 |
proof (rule r) |
154 |
show "Q" |
|
155 |
by (rule mp [OF major \<open>P\<close>]) |
|
156 |
qed |
|
21539 | 157 |
|
158 |
lemma allE: |
|
69590 | 159 |
assumes major: \<open>\<forall>x. P(x)\<close> |
160 |
and r: \<open>P(x) \<Longrightarrow> R\<close> |
|
161 |
shows \<open>R\<close> |
|
71839 | 162 |
proof (rule r) |
163 |
show "P(x)" |
|
164 |
by (rule major [THEN spec]) |
|
165 |
qed |
|
21539 | 166 |
|
69593 | 167 |
text \<open>Duplicates the quantifier; for use with \<^ML>\<open>eresolve_tac\<close>.\<close> |
21539 | 168 |
lemma all_dupE: |
69590 | 169 |
assumes major: \<open>\<forall>x. P(x)\<close> |
170 |
and r: \<open>\<lbrakk>P(x); \<forall>x. P(x)\<rbrakk> \<Longrightarrow> R\<close> |
|
171 |
shows \<open>R\<close> |
|
71839 | 172 |
proof (rule r) |
173 |
show "P(x)" |
|
174 |
by (rule major [THEN spec]) |
|
175 |
qed (rule major) |
|
176 |
||
21539 | 177 |
|
178 |
||
62020 | 179 |
subsubsection \<open>Negation rules, which translate between \<open>\<not> P\<close> and \<open>P \<longrightarrow> False\<close>\<close> |
21539 | 180 |
|
69590 | 181 |
lemma notI: \<open>(P \<Longrightarrow> False) \<Longrightarrow> \<not> P\<close> |
21539 | 182 |
unfolding not_def by (erule impI) |
183 |
||
69590 | 184 |
lemma notE: \<open>\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R\<close> |
21539 | 185 |
unfolding not_def by (erule mp [THEN FalseE]) |
186 |
||
69590 | 187 |
lemma rev_notE: \<open>\<lbrakk>P; \<not> P\<rbrakk> \<Longrightarrow> R\<close> |
21539 | 188 |
by (erule notE) |
189 |
||
62020 | 190 |
text \<open>This is useful with the special implication rules for each kind of \<open>P\<close>.\<close> |
21539 | 191 |
lemma not_to_imp: |
69590 | 192 |
assumes \<open>\<not> P\<close> |
193 |
and r: \<open>P \<longrightarrow> False \<Longrightarrow> Q\<close> |
|
194 |
shows \<open>Q\<close> |
|
21539 | 195 |
apply (rule r) |
196 |
apply (rule impI) |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
197 |
apply (erule notE [OF \<open>\<not> P\<close>]) |
21539 | 198 |
done |
199 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
200 |
text \<open> |
62020 | 201 |
For substitution into an assumption \<open>P\<close>, reduce \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, substitute into this implication, then apply \<open>impI\<close> to |
202 |
move \<open>P\<close> back into the assumptions. |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
203 |
\<close> |
69590 | 204 |
lemma rev_mp: \<open>\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close> |
21539 | 205 |
by (erule mp) |
206 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
207 |
text \<open>Contrapositive of an inference rule.\<close> |
21539 | 208 |
lemma contrapos: |
69590 | 209 |
assumes major: \<open>\<not> Q\<close> |
210 |
and minor: \<open>P \<Longrightarrow> Q\<close> |
|
211 |
shows \<open>\<not> P\<close> |
|
21539 | 212 |
apply (rule major [THEN notE, THEN notI]) |
213 |
apply (erule minor) |
|
214 |
done |
|
215 |
||
216 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
217 |
subsubsection \<open>Modus Ponens Tactics\<close> |
21539 | 218 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
219 |
text \<open> |
62020 | 220 |
Finds \<open>P \<longrightarrow> Q\<close> and P in the assumptions, replaces implication by |
221 |
\<open>Q\<close>. |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
222 |
\<close> |
60770 | 223 |
ML \<open> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
224 |
fun mp_tac ctxt i = |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
225 |
eresolve_tac ctxt @{thms notE impE} i THEN assume_tac ctxt i; |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
226 |
fun eq_mp_tac ctxt i = |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
227 |
eresolve_tac ctxt @{thms notE impE} i THEN eq_assume_tac i; |
60770 | 228 |
\<close> |
21539 | 229 |
|
230 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
231 |
subsection \<open>If-and-only-if\<close> |
21539 | 232 |
|
69590 | 233 |
lemma iffI: \<open>\<lbrakk>P \<Longrightarrow> Q; Q \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> Q\<close> |
71839 | 234 |
unfolding iff_def |
235 |
by (rule conjI; erule impI) |
|
21539 | 236 |
|
237 |
lemma iffE: |
|
69590 | 238 |
assumes major: \<open>P \<longleftrightarrow> Q\<close> |
71839 | 239 |
and r: \<open>\<lbrakk>P \<longrightarrow> Q; Q \<longrightarrow> P\<rbrakk> \<Longrightarrow> R\<close> |
69590 | 240 |
shows \<open>R\<close> |
71839 | 241 |
using major |
242 |
unfolding iff_def |
|
243 |
apply (rule conjE) |
|
21539 | 244 |
apply (erule r) |
245 |
apply assumption |
|
246 |
done |
|
247 |
||
248 |
||
62020 | 249 |
subsubsection \<open>Destruct rules for \<open>\<longleftrightarrow>\<close> similar to Modus Ponens\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
250 |
|
69590 | 251 |
lemma iffD1: \<open>\<lbrakk>P \<longleftrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close> |
71839 | 252 |
unfolding iff_def |
21539 | 253 |
apply (erule conjunct1 [THEN mp]) |
254 |
apply assumption |
|
255 |
done |
|
256 |
||
69590 | 257 |
lemma iffD2: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q\<rbrakk> \<Longrightarrow> P\<close> |
71839 | 258 |
unfolding iff_def |
21539 | 259 |
apply (erule conjunct2 [THEN mp]) |
260 |
apply assumption |
|
261 |
done |
|
262 |
||
69590 | 263 |
lemma rev_iffD1: \<open>\<lbrakk>P; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close> |
21539 | 264 |
apply (erule iffD1) |
265 |
apply assumption |
|
266 |
done |
|
267 |
||
69590 | 268 |
lemma rev_iffD2: \<open>\<lbrakk>Q; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> P\<close> |
21539 | 269 |
apply (erule iffD2) |
270 |
apply assumption |
|
271 |
done |
|
272 |
||
69590 | 273 |
lemma iff_refl: \<open>P \<longleftrightarrow> P\<close> |
21539 | 274 |
by (rule iffI) |
275 |
||
69590 | 276 |
lemma iff_sym: \<open>Q \<longleftrightarrow> P \<Longrightarrow> P \<longleftrightarrow> Q\<close> |
21539 | 277 |
apply (erule iffE) |
278 |
apply (rule iffI) |
|
279 |
apply (assumption | erule mp)+ |
|
280 |
done |
|
281 |
||
69590 | 282 |
lemma iff_trans: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q \<longleftrightarrow> R\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> R\<close> |
21539 | 283 |
apply (rule iffI) |
284 |
apply (assumption | erule iffE | erule (1) notE impE)+ |
|
285 |
done |
|
286 |
||
287 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
288 |
subsection \<open>Unique existence\<close> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
289 |
|
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
290 |
text \<open> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
291 |
NOTE THAT the following 2 quantifications: |
21539 | 292 |
|
63906
fa799a8e4adc
repair LaTeX dropout from f83ef97d8d7d
Lars Hupel <lars.hupel@mytum.de>
parents:
63901
diff
changeset
|
293 |
\<^item> \<open>\<exists>!x\<close> such that [\<open>\<exists>!y\<close> such that P(x,y)] (sequential) |
fa799a8e4adc
repair LaTeX dropout from f83ef97d8d7d
Lars Hupel <lars.hupel@mytum.de>
parents:
63901
diff
changeset
|
294 |
\<^item> \<open>\<exists>!x,y\<close> such that P(x,y) (simultaneous) |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
295 |
|
63906
fa799a8e4adc
repair LaTeX dropout from f83ef97d8d7d
Lars Hupel <lars.hupel@mytum.de>
parents:
63901
diff
changeset
|
296 |
do NOT mean the same thing. The parser treats \<open>\<exists>!x y.P(x,y)\<close> as sequential. |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
297 |
\<close> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
298 |
|
69590 | 299 |
lemma ex1I: \<open>P(a) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> x = a) \<Longrightarrow> \<exists>!x. P(x)\<close> |
71839 | 300 |
unfolding ex1_def |
23393 | 301 |
apply (assumption | rule exI conjI allI impI)+ |
21539 | 302 |
done |
303 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
304 |
text \<open>Sometimes easier to use: the premises have no shared variables. Safe!\<close> |
69590 | 305 |
lemma ex_ex1I: \<open>\<exists>x. P(x) \<Longrightarrow> (\<And>x y. \<lbrakk>P(x); P(y)\<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> \<exists>!x. P(x)\<close> |
23393 | 306 |
apply (erule exE) |
307 |
apply (rule ex1I) |
|
308 |
apply assumption |
|
309 |
apply assumption |
|
21539 | 310 |
done |
311 |
||
69590 | 312 |
lemma ex1E: \<open>\<exists>! x. P(x) \<Longrightarrow> (\<And>x. \<lbrakk>P(x); \<forall>y. P(y) \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R\<close> |
71839 | 313 |
unfolding ex1_def |
21539 | 314 |
apply (assumption | erule exE conjE)+ |
315 |
done |
|
316 |
||
317 |
||
62020 | 318 |
subsubsection \<open>\<open>\<longleftrightarrow>\<close> congruence rules for simplification\<close> |
21539 | 319 |
|
62020 | 320 |
text \<open>Use \<open>iffE\<close> on a premise. For \<open>conj_cong\<close>, \<open>imp_cong\<close>, \<open>all_cong\<close>, \<open>ex_cong\<close>.\<close> |
60770 | 321 |
ML \<open> |
59529 | 322 |
fun iff_tac ctxt prems i = |
323 |
resolve_tac ctxt (prems RL @{thms iffE}) i THEN |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
324 |
REPEAT1 (eresolve_tac ctxt @{thms asm_rl mp} i); |
60770 | 325 |
\<close> |
21539 | 326 |
|
59529 | 327 |
method_setup iff = |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
328 |
\<open>Attrib.thms >> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
329 |
(fn prems => fn ctxt => SIMPLE_METHOD' (iff_tac ctxt prems))\<close> |
59529 | 330 |
|
21539 | 331 |
lemma conj_cong: |
69590 | 332 |
assumes \<open>P \<longleftrightarrow> P'\<close> |
333 |
and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close> |
|
334 |
shows \<open>(P \<and> Q) \<longleftrightarrow> (P' \<and> Q')\<close> |
|
21539 | 335 |
apply (insert assms) |
59529 | 336 |
apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+ |
21539 | 337 |
done |
338 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
339 |
text \<open>Reversed congruence rule! Used in ZF/Order.\<close> |
21539 | 340 |
lemma conj_cong2: |
69590 | 341 |
assumes \<open>P \<longleftrightarrow> P'\<close> |
342 |
and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close> |
|
343 |
shows \<open>(Q \<and> P) \<longleftrightarrow> (Q' \<and> P')\<close> |
|
21539 | 344 |
apply (insert assms) |
59529 | 345 |
apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+ |
21539 | 346 |
done |
347 |
||
348 |
lemma disj_cong: |
|
69590 | 349 |
assumes \<open>P \<longleftrightarrow> P'\<close> and \<open>Q \<longleftrightarrow> Q'\<close> |
350 |
shows \<open>(P \<or> Q) \<longleftrightarrow> (P' \<or> Q')\<close> |
|
21539 | 351 |
apply (insert assms) |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
352 |
apply (erule iffE disjE disjI1 disjI2 | |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
353 |
assumption | rule iffI | erule (1) notE impE)+ |
21539 | 354 |
done |
355 |
||
356 |
lemma imp_cong: |
|
69590 | 357 |
assumes \<open>P \<longleftrightarrow> P'\<close> |
358 |
and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close> |
|
359 |
shows \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (P' \<longrightarrow> Q')\<close> |
|
21539 | 360 |
apply (insert assms) |
59529 | 361 |
apply (assumption | rule iffI impI | erule iffE | erule (1) notE impE | iff assms)+ |
21539 | 362 |
done |
363 |
||
69590 | 364 |
lemma iff_cong: \<open>\<lbrakk>P \<longleftrightarrow> P'; Q \<longleftrightarrow> Q'\<rbrakk> \<Longrightarrow> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P' \<longleftrightarrow> Q')\<close> |
21539 | 365 |
apply (erule iffE | assumption | rule iffI | erule (1) notE impE)+ |
366 |
done |
|
367 |
||
69590 | 368 |
lemma not_cong: \<open>P \<longleftrightarrow> P' \<Longrightarrow> \<not> P \<longleftrightarrow> \<not> P'\<close> |
21539 | 369 |
apply (assumption | rule iffI notI | erule (1) notE impE | erule iffE notE)+ |
370 |
done |
|
371 |
||
372 |
lemma all_cong: |
|
69590 | 373 |
assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close> |
374 |
shows \<open>(\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))\<close> |
|
59529 | 375 |
apply (assumption | rule iffI allI | erule (1) notE impE | erule allE | iff assms)+ |
21539 | 376 |
done |
377 |
||
378 |
lemma ex_cong: |
|
69590 | 379 |
assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close> |
380 |
shows \<open>(\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))\<close> |
|
59529 | 381 |
apply (erule exE | assumption | rule iffI exI | erule (1) notE impE | iff assms)+ |
21539 | 382 |
done |
383 |
||
384 |
lemma ex1_cong: |
|
69590 | 385 |
assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close> |
386 |
shows \<open>(\<exists>!x. P(x)) \<longleftrightarrow> (\<exists>!x. Q(x))\<close> |
|
59529 | 387 |
apply (erule ex1E spec [THEN mp] | assumption | rule iffI ex1I | erule (1) notE impE | iff assms)+ |
21539 | 388 |
done |
389 |
||
390 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
391 |
subsection \<open>Equality rules\<close> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
392 |
|
69590 | 393 |
lemma sym: \<open>a = b \<Longrightarrow> b = a\<close> |
21539 | 394 |
apply (erule subst) |
395 |
apply (rule refl) |
|
396 |
done |
|
397 |
||
69590 | 398 |
lemma trans: \<open>\<lbrakk>a = b; b = c\<rbrakk> \<Longrightarrow> a = c\<close> |
21539 | 399 |
apply (erule subst, assumption) |
400 |
done |
|
401 |
||
69590 | 402 |
lemma not_sym: \<open>b \<noteq> a \<Longrightarrow> a \<noteq> b\<close> |
21539 | 403 |
apply (erule contrapos) |
404 |
apply (erule sym) |
|
405 |
done |
|
406 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
407 |
text \<open> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
408 |
Two theorems for rewriting only one instance of a definition: |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
409 |
the first for definitions of formulae and the second for terms. |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
410 |
\<close> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
411 |
|
69590 | 412 |
lemma def_imp_iff: \<open>(A \<equiv> B) \<Longrightarrow> A \<longleftrightarrow> B\<close> |
21539 | 413 |
apply unfold |
414 |
apply (rule iff_refl) |
|
415 |
done |
|
416 |
||
69590 | 417 |
lemma meta_eq_to_obj_eq: \<open>(A \<equiv> B) \<Longrightarrow> A = B\<close> |
21539 | 418 |
apply unfold |
419 |
apply (rule refl) |
|
420 |
done |
|
421 |
||
69590 | 422 |
lemma meta_eq_to_iff: \<open>x \<equiv> y \<Longrightarrow> x \<longleftrightarrow> y\<close> |
21539 | 423 |
by unfold (rule iff_refl) |
424 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
425 |
text \<open>Substitution.\<close> |
69590 | 426 |
lemma ssubst: \<open>\<lbrakk>b = a; P(a)\<rbrakk> \<Longrightarrow> P(b)\<close> |
21539 | 427 |
apply (drule sym) |
428 |
apply (erule (1) subst) |
|
429 |
done |
|
430 |
||
62020 | 431 |
text \<open>A special case of \<open>ex1E\<close> that would otherwise need quantifier |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
432 |
expansion.\<close> |
69590 | 433 |
lemma ex1_equalsE: \<open>\<lbrakk>\<exists>!x. P(x); P(a); P(b)\<rbrakk> \<Longrightarrow> a = b\<close> |
21539 | 434 |
apply (erule ex1E) |
435 |
apply (rule trans) |
|
436 |
apply (rule_tac [2] sym) |
|
437 |
apply (assumption | erule spec [THEN mp])+ |
|
438 |
done |
|
439 |
||
440 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
441 |
subsection \<open>Simplifications of assumed implications\<close> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
442 |
|
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
443 |
text \<open> |
62020 | 444 |
Roy Dyckhoff has proved that \<open>conj_impE\<close>, \<open>disj_impE\<close>, and |
69593 | 445 |
\<open>imp_impE\<close> used with \<^ML>\<open>mp_tac\<close> (restricted to atomic formulae) is |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
446 |
COMPLETE for intuitionistic propositional logic. |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
447 |
|
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
448 |
See R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
449 |
(preprint, University of St Andrews, 1991). |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
450 |
\<close> |
21539 | 451 |
|
452 |
lemma conj_impE: |
|
69590 | 453 |
assumes major: \<open>(P \<and> Q) \<longrightarrow> S\<close> |
454 |
and r: \<open>P \<longrightarrow> (Q \<longrightarrow> S) \<Longrightarrow> R\<close> |
|
455 |
shows \<open>R\<close> |
|
21539 | 456 |
by (assumption | rule conjI impI major [THEN mp] r)+ |
457 |
||
458 |
lemma disj_impE: |
|
69590 | 459 |
assumes major: \<open>(P \<or> Q) \<longrightarrow> S\<close> |
460 |
and r: \<open>\<lbrakk>P \<longrightarrow> S; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> R\<close> |
|
461 |
shows \<open>R\<close> |
|
21539 | 462 |
by (assumption | rule disjI1 disjI2 impI major [THEN mp] r)+ |
463 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
464 |
text \<open>Simplifies the implication. Classical version is stronger. |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
465 |
Still UNSAFE since Q must be provable -- backtracking needed.\<close> |
21539 | 466 |
lemma imp_impE: |
69590 | 467 |
assumes major: \<open>(P \<longrightarrow> Q) \<longrightarrow> S\<close> |
468 |
and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close> |
|
469 |
and r2: \<open>S \<Longrightarrow> R\<close> |
|
470 |
shows \<open>R\<close> |
|
21539 | 471 |
by (assumption | rule impI major [THEN mp] r1 r2)+ |
472 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
473 |
text \<open>Simplifies the implication. Classical version is stronger. |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
474 |
Still UNSAFE since ~P must be provable -- backtracking needed.\<close> |
69590 | 475 |
lemma not_impE: \<open>\<not> P \<longrightarrow> S \<Longrightarrow> (P \<Longrightarrow> False) \<Longrightarrow> (S \<Longrightarrow> R) \<Longrightarrow> R\<close> |
23393 | 476 |
apply (drule mp) |
71839 | 477 |
apply (rule notI | assumption)+ |
21539 | 478 |
done |
479 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
480 |
text \<open>Simplifies the implication. UNSAFE.\<close> |
21539 | 481 |
lemma iff_impE: |
69590 | 482 |
assumes major: \<open>(P \<longleftrightarrow> Q) \<longrightarrow> S\<close> |
483 |
and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close> |
|
484 |
and r2: \<open>\<lbrakk>Q; P \<longrightarrow> S\<rbrakk> \<Longrightarrow> P\<close> |
|
485 |
and r3: \<open>S \<Longrightarrow> R\<close> |
|
486 |
shows \<open>R\<close> |
|
71839 | 487 |
by (assumption | rule iffI impI major [THEN mp] r1 r2 r3)+ |
21539 | 488 |
|
62020 | 489 |
text \<open>What if \<open>(\<forall>x. \<not> \<not> P(x)) \<longrightarrow> \<not> \<not> (\<forall>x. P(x))\<close> is an assumption? |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
490 |
UNSAFE.\<close> |
21539 | 491 |
lemma all_impE: |
69590 | 492 |
assumes major: \<open>(\<forall>x. P(x)) \<longrightarrow> S\<close> |
493 |
and r1: \<open>\<And>x. P(x)\<close> |
|
494 |
and r2: \<open>S \<Longrightarrow> R\<close> |
|
495 |
shows \<open>R\<close> |
|
71839 | 496 |
by (rule allI impI major [THEN mp] r1 r2)+ |
21539 | 497 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
498 |
text \<open> |
62020 | 499 |
Unsafe: \<open>\<exists>x. P(x)) \<longrightarrow> S\<close> is equivalent |
500 |
to \<open>\<forall>x. P(x) \<longrightarrow> S\<close>.\<close> |
|
21539 | 501 |
lemma ex_impE: |
69590 | 502 |
assumes major: \<open>(\<exists>x. P(x)) \<longrightarrow> S\<close> |
503 |
and r: \<open>P(x) \<longrightarrow> S \<Longrightarrow> R\<close> |
|
504 |
shows \<open>R\<close> |
|
71839 | 505 |
by (assumption | rule exI impI major [THEN mp] r)+ |
21539 | 506 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
507 |
text \<open>Courtesy of Krzysztof Grabczewski.\<close> |
69590 | 508 |
lemma disj_imp_disj: \<open>P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> S) \<Longrightarrow> R \<or> S\<close> |
23393 | 509 |
apply (erule disjE) |
21539 | 510 |
apply (rule disjI1) apply assumption |
511 |
apply (rule disjI2) apply assumption |
|
512 |
done |
|
11734 | 513 |
|
60770 | 514 |
ML \<open> |
32172 | 515 |
structure Project_Rule = Project_Rule |
516 |
( |
|
22139 | 517 |
val conjunct1 = @{thm conjunct1} |
518 |
val conjunct2 = @{thm conjunct2} |
|
519 |
val mp = @{thm mp} |
|
32172 | 520 |
) |
60770 | 521 |
\<close> |
18481 | 522 |
|
69605 | 523 |
ML_file \<open>fologic.ML\<close> |
21539 | 524 |
|
69590 | 525 |
lemma thin_refl: \<open>\<lbrakk>x = x; PROP W\<rbrakk> \<Longrightarrow> PROP W\<close> . |
21539 | 526 |
|
60770 | 527 |
ML \<open> |
42799 | 528 |
structure Hypsubst = Hypsubst |
529 |
( |
|
530 |
val dest_eq = FOLogic.dest_eq |
|
74342 | 531 |
val dest_Trueprop = \<^dest_judgment> |
42799 | 532 |
val dest_imp = FOLogic.dest_imp |
533 |
val eq_reflection = @{thm eq_reflection} |
|
534 |
val rev_eq_reflection = @{thm meta_eq_to_obj_eq} |
|
535 |
val imp_intr = @{thm impI} |
|
536 |
val rev_mp = @{thm rev_mp} |
|
537 |
val subst = @{thm subst} |
|
538 |
val sym = @{thm sym} |
|
539 |
val thin_refl = @{thm thin_refl} |
|
540 |
); |
|
541 |
open Hypsubst; |
|
60770 | 542 |
\<close> |
42799 | 543 |
|
69605 | 544 |
ML_file \<open>intprover.ML\<close> |
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
545 |
|
4092 | 546 |
|
60770 | 547 |
subsection \<open>Intuitionistic Reasoning\<close> |
12368 | 548 |
|
69593 | 549 |
setup \<open>Intuitionistic.method_setup \<^binding>\<open>iprover\<close>\<close> |
30165
6ee87f67d9cd
moved generic intuitionistic prover to src/Tools/intuitionistic.ML;
wenzelm
parents:
30160
diff
changeset
|
550 |
|
12349 | 551 |
lemma impE': |
69590 | 552 |
assumes 1: \<open>P \<longrightarrow> Q\<close> |
553 |
and 2: \<open>Q \<Longrightarrow> R\<close> |
|
554 |
and 3: \<open>P \<longrightarrow> Q \<Longrightarrow> P\<close> |
|
555 |
shows \<open>R\<close> |
|
12349 | 556 |
proof - |
69590 | 557 |
from 3 and 1 have \<open>P\<close> . |
558 |
with 1 have \<open>Q\<close> by (rule impE) |
|
559 |
with 2 show \<open>R\<close> . |
|
12349 | 560 |
qed |
561 |
||
562 |
lemma allE': |
|
69590 | 563 |
assumes 1: \<open>\<forall>x. P(x)\<close> |
564 |
and 2: \<open>P(x) \<Longrightarrow> \<forall>x. P(x) \<Longrightarrow> Q\<close> |
|
565 |
shows \<open>Q\<close> |
|
12349 | 566 |
proof - |
69590 | 567 |
from 1 have \<open>P(x)\<close> by (rule spec) |
568 |
from this and 1 show \<open>Q\<close> by (rule 2) |
|
12349 | 569 |
qed |
570 |
||
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12875
diff
changeset
|
571 |
lemma notE': |
69590 | 572 |
assumes 1: \<open>\<not> P\<close> |
573 |
and 2: \<open>\<not> P \<Longrightarrow> P\<close> |
|
574 |
shows \<open>R\<close> |
|
12349 | 575 |
proof - |
69590 | 576 |
from 2 and 1 have \<open>P\<close> . |
577 |
with 1 show \<open>R\<close> by (rule notE) |
|
12349 | 578 |
qed |
579 |
||
580 |
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE |
|
581 |
and [Pure.intro!] = iffI conjI impI TrueI notI allI refl |
|
582 |
and [Pure.elim 2] = allE notE' impE' |
|
583 |
and [Pure.intro] = exI disjI2 disjI1 |
|
584 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
585 |
setup \<open> |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
586 |
Context_Rules.addSWrapper |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
587 |
(fn ctxt => fn tac => hyp_subst_tac ctxt ORELSE' tac) |
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
588 |
\<close> |
12349 | 589 |
|
590 |
||
69590 | 591 |
lemma iff_not_sym: \<open>\<not> (Q \<longleftrightarrow> P) \<Longrightarrow> \<not> (P \<longleftrightarrow> Q)\<close> |
17591 | 592 |
by iprover |
12368 | 593 |
|
594 |
lemmas [sym] = sym iff_sym not_sym iff_not_sym |
|
595 |
and [Pure.elim?] = iffD1 iffD2 impE |
|
596 |
||
597 |
||
69590 | 598 |
lemma eq_commute: \<open>a = b \<longleftrightarrow> b = a\<close> |
71839 | 599 |
by iprover |
600 |
||
601 |
||
602 |
subsection \<open>Polymorphic congruence rules\<close> |
|
603 |
||
604 |
lemma subst_context: \<open>a = b \<Longrightarrow> t(a) = t(b)\<close> |
|
605 |
by iprover |
|
606 |
||
607 |
lemma subst_context2: \<open>\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> t(a,c) = t(b,d)\<close> |
|
608 |
by iprover |
|
609 |
||
610 |
lemma subst_context3: \<open>\<lbrakk>a = b; c = d; e = f\<rbrakk> \<Longrightarrow> t(a,c,e) = t(b,d,f)\<close> |
|
611 |
by iprover |
|
612 |
||
613 |
text \<open> |
|
614 |
Useful with \<^ML>\<open>eresolve_tac\<close> for proving equalities from known |
|
615 |
equalities. |
|
616 |
||
617 |
a = b |
|
618 |
| | |
|
619 |
c = d |
|
620 |
\<close> |
|
621 |
lemma box_equals: \<open>\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d\<close> |
|
622 |
by iprover |
|
623 |
||
624 |
text \<open>Dual of \<open>box_equals\<close>: for proving equalities backwards.\<close> |
|
625 |
lemma simp_equals: \<open>\<lbrakk>a = c; b = d; c = d\<rbrakk> \<Longrightarrow> a = b\<close> |
|
626 |
by iprover |
|
627 |
||
628 |
||
629 |
subsubsection \<open>Congruence rules for predicate letters\<close> |
|
630 |
||
631 |
lemma pred1_cong: \<open>a = a' \<Longrightarrow> P(a) \<longleftrightarrow> P(a')\<close> |
|
632 |
by iprover |
|
633 |
||
634 |
lemma pred2_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> P(a,b) \<longleftrightarrow> P(a',b')\<close> |
|
635 |
by iprover |
|
636 |
||
637 |
lemma pred3_cong: \<open>\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> P(a,b,c) \<longleftrightarrow> P(a',b',c')\<close> |
|
638 |
by iprover |
|
639 |
||
640 |
text \<open>Special case for the equality predicate!\<close> |
|
641 |
lemma eq_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> a = b \<longleftrightarrow> a' = b'\<close> |
|
642 |
by iprover |
|
13435 | 643 |
|
644 |
||
60770 | 645 |
subsection \<open>Atomizing meta-level rules\<close> |
11677 | 646 |
|
69590 | 647 |
lemma atomize_all [atomize]: \<open>(\<And>x. P(x)) \<equiv> Trueprop (\<forall>x. P(x))\<close> |
11976 | 648 |
proof |
69590 | 649 |
assume \<open>\<And>x. P(x)\<close> |
650 |
then show \<open>\<forall>x. P(x)\<close> .. |
|
11677 | 651 |
next |
69590 | 652 |
assume \<open>\<forall>x. P(x)\<close> |
653 |
then show \<open>\<And>x. P(x)\<close> .. |
|
11677 | 654 |
qed |
655 |
||
69590 | 656 |
lemma atomize_imp [atomize]: \<open>(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)\<close> |
11976 | 657 |
proof |
69590 | 658 |
assume \<open>A \<Longrightarrow> B\<close> |
659 |
then show \<open>A \<longrightarrow> B\<close> .. |
|
11677 | 660 |
next |
69590 | 661 |
assume \<open>A \<longrightarrow> B\<close> and \<open>A\<close> |
662 |
then show \<open>B\<close> by (rule mp) |
|
11677 | 663 |
qed |
664 |
||
69590 | 665 |
lemma atomize_eq [atomize]: \<open>(x \<equiv> y) \<equiv> Trueprop (x = y)\<close> |
11976 | 666 |
proof |
69590 | 667 |
assume \<open>x \<equiv> y\<close> |
668 |
show \<open>x = y\<close> unfolding \<open>x \<equiv> y\<close> by (rule refl) |
|
11677 | 669 |
next |
69590 | 670 |
assume \<open>x = y\<close> |
671 |
then show \<open>x \<equiv> y\<close> by (rule eq_reflection) |
|
11677 | 672 |
qed |
673 |
||
69590 | 674 |
lemma atomize_iff [atomize]: \<open>(A \<equiv> B) \<equiv> Trueprop (A \<longleftrightarrow> B)\<close> |
18813 | 675 |
proof |
69590 | 676 |
assume \<open>A \<equiv> B\<close> |
677 |
show \<open>A \<longleftrightarrow> B\<close> unfolding \<open>A \<equiv> B\<close> by (rule iff_refl) |
|
18813 | 678 |
next |
69590 | 679 |
assume \<open>A \<longleftrightarrow> B\<close> |
680 |
then show \<open>A \<equiv> B\<close> by (rule iff_reflection) |
|
18813 | 681 |
qed |
682 |
||
69590 | 683 |
lemma atomize_conj [atomize]: \<open>(A &&& B) \<equiv> Trueprop (A \<and> B)\<close> |
11976 | 684 |
proof |
69590 | 685 |
assume conj: \<open>A &&& B\<close> |
686 |
show \<open>A \<and> B\<close> |
|
19120
353d349740de
not_equal: replaced syntax translation by abbreviation;
wenzelm
parents:
18861
diff
changeset
|
687 |
proof (rule conjI) |
69590 | 688 |
from conj show \<open>A\<close> by (rule conjunctionD1) |
689 |
from conj show \<open>B\<close> by (rule conjunctionD2) |
|
19120
353d349740de
not_equal: replaced syntax translation by abbreviation;
wenzelm
parents:
18861
diff
changeset
|
690 |
qed |
11953 | 691 |
next |
69590 | 692 |
assume conj: \<open>A \<and> B\<close> |
693 |
show \<open>A &&& B\<close> |
|
19120
353d349740de
not_equal: replaced syntax translation by abbreviation;
wenzelm
parents:
18861
diff
changeset
|
694 |
proof - |
69590 | 695 |
from conj show \<open>A\<close> .. |
696 |
from conj show \<open>B\<close> .. |
|
11953 | 697 |
qed |
698 |
qed |
|
699 |
||
12368 | 700 |
lemmas [symmetric, rulify] = atomize_all atomize_imp |
18861 | 701 |
and [symmetric, defn] = atomize_all atomize_imp atomize_eq atomize_iff |
11771 | 702 |
|
11848 | 703 |
|
60770 | 704 |
subsection \<open>Atomizing elimination rules\<close> |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26286
diff
changeset
|
705 |
|
69590 | 706 |
lemma atomize_exL[atomize_elim]: \<open>(\<And>x. P(x) \<Longrightarrow> Q) \<equiv> ((\<exists>x. P(x)) \<Longrightarrow> Q)\<close> |
57948 | 707 |
by rule iprover+ |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26286
diff
changeset
|
708 |
|
69590 | 709 |
lemma atomize_conjL[atomize_elim]: \<open>(A \<Longrightarrow> B \<Longrightarrow> C) \<equiv> (A \<and> B \<Longrightarrow> C)\<close> |
57948 | 710 |
by rule iprover+ |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26286
diff
changeset
|
711 |
|
69590 | 712 |
lemma atomize_disjL[atomize_elim]: \<open>((A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C) \<equiv> ((A \<or> B \<Longrightarrow> C) \<Longrightarrow> C)\<close> |
57948 | 713 |
by rule iprover+ |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26286
diff
changeset
|
714 |
|
69590 | 715 |
lemma atomize_elimL[atomize_elim]: \<open>(\<And>B. (A \<Longrightarrow> B) \<Longrightarrow> B) \<equiv> Trueprop(A)\<close> .. |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26286
diff
changeset
|
716 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26286
diff
changeset
|
717 |
|
60770 | 718 |
subsection \<open>Calculational rules\<close> |
11848 | 719 |
|
69590 | 720 |
lemma forw_subst: \<open>a = b \<Longrightarrow> P(b) \<Longrightarrow> P(a)\<close> |
11848 | 721 |
by (rule ssubst) |
722 |
||
69590 | 723 |
lemma back_subst: \<open>P(a) \<Longrightarrow> a = b \<Longrightarrow> P(b)\<close> |
11848 | 724 |
by (rule subst) |
725 |
||
60770 | 726 |
text \<open> |
11848 | 727 |
Note that this list of rules is in reverse order of priorities. |
60770 | 728 |
\<close> |
11848 | 729 |
|
12019 | 730 |
lemmas basic_trans_rules [trans] = |
11848 | 731 |
forw_subst |
732 |
back_subst |
|
733 |
rev_mp |
|
734 |
mp |
|
735 |
trans |
|
736 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
737 |
|
60770 | 738 |
subsection \<open>``Let'' declarations\<close> |
13779 | 739 |
|
41229
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents:
39557
diff
changeset
|
740 |
nonterminal letbinds and letbind |
13779 | 741 |
|
69590 | 742 |
definition Let :: \<open>['a::{}, 'a => 'b] \<Rightarrow> ('b::{})\<close> |
743 |
where \<open>Let(s, f) \<equiv> f(s)\<close> |
|
13779 | 744 |
|
745 |
syntax |
|
80917 | 746 |
"_bind" :: \<open>[pttrn, 'a] => letbind\<close> (\<open>(\<open>indent=2 notation=\<open>infix letbind\<close>\<close>_ =/ _)\<close> 10) |
69590 | 747 |
"" :: \<open>letbind => letbinds\<close> (\<open>_\<close>) |
748 |
"_binds" :: \<open>[letbind, letbinds] => letbinds\<close> (\<open>_;/ _\<close>) |
|
80917 | 749 |
"_Let" :: \<open>[letbinds, 'a] => 'a\<close> (\<open>(\<open>notation=\<open>mixfix let in\<close>\<close>let (_)/ in (_))\<close> 10) |
81011
6d34c2bedaa3
clarified markup: avoid conflict of "notation" with "entity", e.g. in "[x,y,z]" without spaces;
wenzelm
parents:
80923
diff
changeset
|
750 |
syntax_consts |
6d34c2bedaa3
clarified markup: avoid conflict of "notation" with "entity", e.g. in "[x,y,z]" without spaces;
wenzelm
parents:
80923
diff
changeset
|
751 |
Let |
13779 | 752 |
translations |
753 |
"_Let(_binds(b, bs), e)" == "_Let(b, _Let(bs, e))" |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
754 |
"let x = a in e" == "CONST Let(a, \<lambda>x. e)" |
13779 | 755 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
756 |
lemma LetI: |
69590 | 757 |
assumes \<open>\<And>x. x = t \<Longrightarrow> P(u(x))\<close> |
758 |
shows \<open>P(let x = t in u(x))\<close> |
|
71839 | 759 |
unfolding Let_def |
21539 | 760 |
apply (rule refl [THEN assms]) |
761 |
done |
|
762 |
||
763 |
||
60770 | 764 |
subsection \<open>Intuitionistic simplification rules\<close> |
26286 | 765 |
|
766 |
lemma conj_simps: |
|
69590 | 767 |
\<open>P \<and> True \<longleftrightarrow> P\<close> |
768 |
\<open>True \<and> P \<longleftrightarrow> P\<close> |
|
769 |
\<open>P \<and> False \<longleftrightarrow> False\<close> |
|
770 |
\<open>False \<and> P \<longleftrightarrow> False\<close> |
|
771 |
\<open>P \<and> P \<longleftrightarrow> P\<close> |
|
772 |
\<open>P \<and> P \<and> Q \<longleftrightarrow> P \<and> Q\<close> |
|
773 |
\<open>P \<and> \<not> P \<longleftrightarrow> False\<close> |
|
774 |
\<open>\<not> P \<and> P \<longleftrightarrow> False\<close> |
|
775 |
\<open>(P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)\<close> |
|
26286 | 776 |
by iprover+ |
777 |
||
778 |
lemma disj_simps: |
|
69590 | 779 |
\<open>P \<or> True \<longleftrightarrow> True\<close> |
780 |
\<open>True \<or> P \<longleftrightarrow> True\<close> |
|
781 |
\<open>P \<or> False \<longleftrightarrow> P\<close> |
|
782 |
\<open>False \<or> P \<longleftrightarrow> P\<close> |
|
783 |
\<open>P \<or> P \<longleftrightarrow> P\<close> |
|
784 |
\<open>P \<or> P \<or> Q \<longleftrightarrow> P \<or> Q\<close> |
|
785 |
\<open>(P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)\<close> |
|
26286 | 786 |
by iprover+ |
787 |
||
788 |
lemma not_simps: |
|
69590 | 789 |
\<open>\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q\<close> |
790 |
\<open>\<not> False \<longleftrightarrow> True\<close> |
|
791 |
\<open>\<not> True \<longleftrightarrow> False\<close> |
|
26286 | 792 |
by iprover+ |
793 |
||
794 |
lemma imp_simps: |
|
69590 | 795 |
\<open>(P \<longrightarrow> False) \<longleftrightarrow> \<not> P\<close> |
796 |
\<open>(P \<longrightarrow> True) \<longleftrightarrow> True\<close> |
|
797 |
\<open>(False \<longrightarrow> P) \<longleftrightarrow> True\<close> |
|
798 |
\<open>(True \<longrightarrow> P) \<longleftrightarrow> P\<close> |
|
799 |
\<open>(P \<longrightarrow> P) \<longleftrightarrow> True\<close> |
|
800 |
\<open>(P \<longrightarrow> \<not> P) \<longleftrightarrow> \<not> P\<close> |
|
26286 | 801 |
by iprover+ |
802 |
||
803 |
lemma iff_simps: |
|
69590 | 804 |
\<open>(True \<longleftrightarrow> P) \<longleftrightarrow> P\<close> |
805 |
\<open>(P \<longleftrightarrow> True) \<longleftrightarrow> P\<close> |
|
806 |
\<open>(P \<longleftrightarrow> P) \<longleftrightarrow> True\<close> |
|
807 |
\<open>(False \<longleftrightarrow> P) \<longleftrightarrow> \<not> P\<close> |
|
808 |
\<open>(P \<longleftrightarrow> False) \<longleftrightarrow> \<not> P\<close> |
|
26286 | 809 |
by iprover+ |
810 |
||
62020 | 811 |
text \<open>The \<open>x = t\<close> versions are needed for the simplification |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
812 |
procedures.\<close> |
26286 | 813 |
lemma quant_simps: |
69590 | 814 |
\<open>\<And>P. (\<forall>x. P) \<longleftrightarrow> P\<close> |
815 |
\<open>(\<forall>x. x = t \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close> |
|
816 |
\<open>(\<forall>x. t = x \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close> |
|
817 |
\<open>\<And>P. (\<exists>x. P) \<longleftrightarrow> P\<close> |
|
818 |
\<open>\<exists>x. x = t\<close> |
|
819 |
\<open>\<exists>x. t = x\<close> |
|
820 |
\<open>(\<exists>x. x = t \<and> P(x)) \<longleftrightarrow> P(t)\<close> |
|
821 |
\<open>(\<exists>x. t = x \<and> P(x)) \<longleftrightarrow> P(t)\<close> |
|
26286 | 822 |
by iprover+ |
823 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
824 |
text \<open>These are NOT supplied by default!\<close> |
26286 | 825 |
lemma distrib_simps: |
69590 | 826 |
\<open>P \<and> (Q \<or> R) \<longleftrightarrow> P \<and> Q \<or> P \<and> R\<close> |
827 |
\<open>(Q \<or> R) \<and> P \<longleftrightarrow> Q \<and> P \<or> R \<and> P\<close> |
|
828 |
\<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close> |
|
26286 | 829 |
by iprover+ |
830 |
||
71919 | 831 |
lemma subst_all: |
832 |
\<open>(\<And>x. x = a \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close> |
|
833 |
\<open>(\<And>x. a = x \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close> |
|
71959 | 834 |
proof - |
835 |
show \<open>(\<And>x. x = a \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close> |
|
836 |
proof (rule equal_intr_rule) |
|
837 |
assume *: \<open>\<And>x. x = a \<Longrightarrow> PROP P(x)\<close> |
|
838 |
show \<open>PROP P(a)\<close> |
|
839 |
by (rule *) (rule refl) |
|
840 |
next |
|
841 |
fix x |
|
842 |
assume \<open>PROP P(a)\<close> and \<open>x = a\<close> |
|
843 |
from \<open>x = a\<close> have \<open>x \<equiv> a\<close> |
|
844 |
by (rule eq_reflection) |
|
845 |
with \<open>PROP P(a)\<close> show \<open>PROP P(x)\<close> |
|
846 |
by simp |
|
847 |
qed |
|
848 |
show \<open>(\<And>x. a = x \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close> |
|
849 |
proof (rule equal_intr_rule) |
|
850 |
assume *: \<open>\<And>x. a = x \<Longrightarrow> PROP P(x)\<close> |
|
851 |
show \<open>PROP P(a)\<close> |
|
852 |
by (rule *) (rule refl) |
|
853 |
next |
|
854 |
fix x |
|
855 |
assume \<open>PROP P(a)\<close> and \<open>a = x\<close> |
|
856 |
from \<open>a = x\<close> have \<open>a \<equiv> x\<close> |
|
857 |
by (rule eq_reflection) |
|
858 |
with \<open>PROP P(a)\<close> show \<open>PROP P(x)\<close> |
|
859 |
by simp |
|
860 |
qed |
|
71919 | 861 |
qed |
862 |
||
26286 | 863 |
|
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
864 |
subsubsection \<open>Conversion into rewrite rules\<close> |
26286 | 865 |
|
69590 | 866 |
lemma P_iff_F: \<open>\<not> P \<Longrightarrow> (P \<longleftrightarrow> False)\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
867 |
by iprover |
69590 | 868 |
lemma iff_reflection_F: \<open>\<not> P \<Longrightarrow> (P \<equiv> False)\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
869 |
by (rule P_iff_F [THEN iff_reflection]) |
26286 | 870 |
|
69590 | 871 |
lemma P_iff_T: \<open>P \<Longrightarrow> (P \<longleftrightarrow> True)\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
872 |
by iprover |
69590 | 873 |
lemma iff_reflection_T: \<open>P \<Longrightarrow> (P \<equiv> True)\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
874 |
by (rule P_iff_T [THEN iff_reflection]) |
26286 | 875 |
|
876 |
||
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
877 |
subsubsection \<open>More rewrite rules\<close> |
26286 | 878 |
|
69590 | 879 |
lemma conj_commute: \<open>P \<and> Q \<longleftrightarrow> Q \<and> P\<close> by iprover |
880 |
lemma conj_left_commute: \<open>P \<and> (Q \<and> R) \<longleftrightarrow> Q \<and> (P \<and> R)\<close> by iprover |
|
26286 | 881 |
lemmas conj_comms = conj_commute conj_left_commute |
882 |
||
69590 | 883 |
lemma disj_commute: \<open>P \<or> Q \<longleftrightarrow> Q \<or> P\<close> by iprover |
884 |
lemma disj_left_commute: \<open>P \<or> (Q \<or> R) \<longleftrightarrow> Q \<or> (P \<or> R)\<close> by iprover |
|
26286 | 885 |
lemmas disj_comms = disj_commute disj_left_commute |
886 |
||
69590 | 887 |
lemma conj_disj_distribL: \<open>P \<and> (Q \<or> R) \<longleftrightarrow> (P \<and> Q \<or> P \<and> R)\<close> by iprover |
888 |
lemma conj_disj_distribR: \<open>(P \<or> Q) \<and> R \<longleftrightarrow> (P \<and> R \<or> Q \<and> R)\<close> by iprover |
|
26286 | 889 |
|
69590 | 890 |
lemma disj_conj_distribL: \<open>P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)\<close> by iprover |
891 |
lemma disj_conj_distribR: \<open>(P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close> by iprover |
|
26286 | 892 |
|
69590 | 893 |
lemma imp_conj_distrib: \<open>(P \<longrightarrow> (Q \<and> R)) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close> by iprover |
894 |
lemma imp_conj: \<open>((P \<and> Q) \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close> by iprover |
|
895 |
lemma imp_disj: \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close> by iprover |
|
26286 | 896 |
|
69590 | 897 |
lemma de_Morgan_disj: \<open>(\<not> (P \<or> Q)) \<longleftrightarrow> (\<not> P \<and> \<not> Q)\<close> by iprover |
26286 | 898 |
|
69590 | 899 |
lemma not_ex: \<open>(\<not> (\<exists>x. P(x))) \<longleftrightarrow> (\<forall>x. \<not> P(x))\<close> by iprover |
900 |
lemma imp_ex: \<open>((\<exists>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> Q)\<close> by iprover |
|
26286 | 901 |
|
69590 | 902 |
lemma ex_disj_distrib: \<open>(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> ((\<exists>x. P(x)) \<or> (\<exists>x. Q(x)))\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
903 |
by iprover |
26286 | 904 |
|
69590 | 905 |
lemma all_conj_distrib: \<open>(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x. P(x)) \<and> (\<forall>x. Q(x)))\<close> |
61487
f8cb97e0fd0b
more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents:
61378
diff
changeset
|
906 |
by iprover |
26286 | 907 |
|
4854 | 908 |
end |