equal
deleted
inserted
replaced
2 |
2 |
3 lemma "((\<exists>x. \<forall>y. p(x)=p(y)) = ((\<exists>x. q(x))=(\<forall>y. p(y)))) = |
3 lemma "((\<exists>x. \<forall>y. p(x)=p(y)) = ((\<exists>x. q(x))=(\<forall>y. p(y)))) = |
4 ((\<exists>x. \<forall>y. q(x)=q(y)) = ((\<exists>x. p(x))=(\<forall>y. q(y))))" |
4 ((\<exists>x. \<forall>y. q(x)=q(y)) = ((\<exists>x. p(x))=(\<forall>y. q(y))))" |
5 by blast |
5 by blast |
6 |
6 |
7 text{*\noindent Until now, we have proved everything using only induction and |
7 text\<open>\noindent Until now, we have proved everything using only induction and |
8 simplification. Substantial proofs require more elaborate types of |
8 simplification. Substantial proofs require more elaborate types of |
9 inference.*} |
9 inference.\<close> |
10 |
10 |
11 lemma "(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and> |
11 lemma "(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and> |
12 \<not> (\<exists>x. grocer(x) \<and> healthy(x)) \<and> |
12 \<not> (\<exists>x. grocer(x) \<and> healthy(x)) \<and> |
13 (\<forall>x. industrious(x) \<and> grocer(x) \<longrightarrow> honest(x)) \<and> |
13 (\<forall>x. industrious(x) \<and> grocer(x) \<longrightarrow> honest(x)) \<and> |
14 (\<forall>x. cyclist(x) \<longrightarrow> industrious(x)) \<and> |
14 (\<forall>x. cyclist(x) \<longrightarrow> industrious(x)) \<and> |
18 |
18 |
19 lemma "(\<Union>i\<in>I. A(i)) \<inter> (\<Union>j\<in>J. B(j)) = |
19 lemma "(\<Union>i\<in>I. A(i)) \<inter> (\<Union>j\<in>J. B(j)) = |
20 (\<Union>i\<in>I. \<Union>j\<in>J. A(i) \<inter> B(j))" |
20 (\<Union>i\<in>I. \<Union>j\<in>J. A(i) \<inter> B(j))" |
21 by blast |
21 by blast |
22 |
22 |
23 text {* |
23 text \<open> |
24 @{thm[display] mult_is_0} |
24 @{thm[display] mult_is_0} |
25 \rulename{mult_is_0}} |
25 \rulename{mult_is_0}} |
26 |
26 |
27 @{thm[display] finite_Un} |
27 @{thm[display] finite_Un} |
28 \rulename{finite_Un}} |
28 \rulename{finite_Un}} |
29 *} |
29 \<close> |
30 |
30 |
31 |
31 |
32 lemma [iff]: "(xs@ys = []) = (xs=[] & ys=[])" |
32 lemma [iff]: "(xs@ys = []) = (xs=[] & ys=[])" |
33 apply (induct_tac xs) |
33 apply (induct_tac xs) |
34 by (simp_all) |
34 by (simp_all) |