|
1 (*<*) |
1 theory Rules = Main: |
2 theory Rules = Main: |
|
3 (*>*) |
2 |
4 |
3 section{*The Rules of the Game*} |
5 section{*The Rules of the Game*} |
4 |
6 |
5 |
7 |
6 subsection{*Introduction Rules*} |
8 subsection{*Introduction Rules*} |
7 |
9 |
|
10 text{* Introduction rules for propositional logic: |
|
11 \begin{center} |
|
12 \begin{tabular}{ll} |
|
13 @{thm[source]impI} & @{thm impI[no_vars]} \\ |
|
14 @{thm[source]conjI} & @{thm conjI[no_vars]} \\ |
|
15 @{thm[source]disjI1} & @{thm disjI1[no_vars]} \\ |
|
16 @{thm[source]disjI2} & @{thm disjI2[no_vars]} \\ |
|
17 @{thm[source]iffI} & @{thm iffI[no_vars]} |
|
18 \end{tabular} |
|
19 \end{center} |
|
20 *} |
|
21 |
|
22 (*<*) |
8 thm impI conjI disjI1 disjI2 iffI |
23 thm impI conjI disjI1 disjI2 iffI |
|
24 (*>*) |
9 |
25 |
10 lemma "A \<Longrightarrow> B \<longrightarrow> A \<and> (B \<and> A)" |
26 lemma "A \<Longrightarrow> B \<longrightarrow> A \<and> (B \<and> A)" |
11 apply(rule impI) |
27 apply(rule impI) |
12 apply(rule conjI) |
28 apply(rule conjI) |
13 apply assumption |
29 apply assumption |
17 done |
33 done |
18 |
34 |
19 |
35 |
20 subsection{*Elimination Rules*} |
36 subsection{*Elimination Rules*} |
21 |
37 |
|
38 text{* Elimination rules for propositional logic: |
|
39 \begin{center} |
|
40 \begin{tabular}{ll} |
|
41 @{thm[source]impE} & @{thm impE[no_vars]} \\ |
|
42 @{thm[source]mp} & @{thm mp[no_vars]} \\ |
|
43 @{thm[source]conjE} & @{thm conjE[no_vars]} \\ |
|
44 @{thm[source]disjE} & @{thm disjE[no_vars]} |
|
45 \end{tabular} |
|
46 \end{center} |
|
47 *} |
|
48 |
|
49 (*<*) |
22 thm impE mp |
50 thm impE mp |
23 thm conjE |
51 thm conjE |
24 thm disjE |
52 thm disjE |
|
53 (*>*) |
25 |
54 |
26 lemma disj_swap: "P | Q \<Longrightarrow> Q | P" |
55 lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P" |
27 apply (erule disjE) |
56 apply (erule disjE) |
28 apply (rule disjI2) |
57 apply (rule disjI2) |
29 apply assumption |
58 apply assumption |
30 apply (rule disjI1) |
59 apply (rule disjI1) |
31 apply assumption |
60 apply assumption |
32 done |
61 done |
33 |
62 |
34 |
63 |
35 subsection{*Destruction Rules*} |
64 subsection{*Destruction Rules*} |
36 |
65 |
|
66 text{* Destruction rules for propositional logic: |
|
67 \begin{center} |
|
68 \begin{tabular}{ll} |
|
69 @{thm[source]conjunct1} & @{thm conjunct1[no_vars]} \\ |
|
70 @{thm[source]conjunct2} & @{thm conjunct2[no_vars]} \\ |
|
71 @{thm[source]iffD1} & @{thm iffD1[no_vars]} \\ |
|
72 @{thm[source]iffD2} & @{thm iffD2[no_vars]} |
|
73 \end{tabular} |
|
74 \end{center} |
|
75 *} |
|
76 |
|
77 (*<*) |
37 thm conjunct1 conjunct1 iffD1 iffD2 |
78 thm conjunct1 conjunct1 iffD1 iffD2 |
|
79 (*>*) |
|
80 |
38 lemma conj_swap: "P \<and> Q \<Longrightarrow> Q \<and> P" |
81 lemma conj_swap: "P \<and> Q \<Longrightarrow> Q \<and> P" |
39 apply (rule conjI) |
82 apply (rule conjI) |
40 apply (drule conjunct2) |
83 apply (drule conjunct2) |
41 apply assumption |
84 apply assumption |
42 apply (drule conjunct1) |
85 apply (drule conjunct1) |
44 done |
87 done |
45 |
88 |
46 |
89 |
47 subsection{*Quantifiers*} |
90 subsection{*Quantifiers*} |
48 |
91 |
|
92 text{* Quantifier rules: |
|
93 \begin{center} |
|
94 \begin{tabular}{ll} |
|
95 @{thm[source]allI} & @{thm allI[no_vars]} \\ |
|
96 @{thm[source]exI} & @{thm exI[no_vars]} \\ |
|
97 @{thm[source]allE} & @{thm allE[no_vars]} \\ |
|
98 @{thm[source]exE} & @{thm exE[no_vars]} \\ |
|
99 @{thm[source]spec} & @{thm spec[no_vars]} |
|
100 \end{tabular} |
|
101 \end{center} |
|
102 *} |
49 |
103 |
|
104 (*<*) |
50 thm allI exI |
105 thm allI exI |
51 thm allE exE |
106 thm allE exE |
52 thm spec |
107 thm spec |
|
108 (*>*) |
53 |
109 |
54 lemma "\<exists>x.\<forall>y. P x y \<Longrightarrow> \<forall>y.\<exists>x. P x y" |
110 lemma "\<exists>x.\<forall>y. P x y \<Longrightarrow> \<forall>y.\<exists>x. P x y" |
55 oops |
111 (*<*)oops(*>*) |
56 |
112 |
57 subsection{*Instantiation*} |
113 subsection{*Instantiation*} |
58 |
114 |
59 lemma "\<exists>xs. xs @ xs = xs" |
115 lemma "\<exists>xs. xs @ xs = xs" |
60 apply(rule_tac x = "[]" in exI) |
116 apply(rule_tac x = "[]" in exI) |
61 by simp |
117 by simp |
62 |
118 |
63 lemma "\<forall>xs. f(xs @ xs) = xs \<Longrightarrow> f [] = []" |
119 lemma "\<forall>xs. f(xs @ xs) = xs \<Longrightarrow> f [] = []" |
64 apply(erule_tac x = "[]" in allE) |
120 apply(erule_tac x = "[]" in allE) |
65 by simp |
121 by simp |
66 |
|
67 |
|
68 subsection{*Hilbert's epsilon Operator*} |
|
69 |
|
70 theorem axiom_of_choice: |
|
71 "\<forall>x. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)" |
|
72 by (fast elim!: someI) |
|
73 |
122 |
74 |
123 |
75 subsection{*Automation*} |
124 subsection{*Automation*} |
76 |
125 |
77 lemma "(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and> |
126 lemma "(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and> |
91 by blast |
140 by blast |
92 |
141 |
93 |
142 |
94 subsection{*Forward Proof*} |
143 subsection{*Forward Proof*} |
95 |
144 |
96 thm rev_rev_ident |
145 text{* |
97 thm rev_rev_ident[of "[]"] |
146 Instantiation of unknowns (in left-to-right order): |
|
147 \begin{center} |
|
148 \begin{tabular}{@ {}ll@ {}} |
|
149 @{thm[source]append_self_conv} & @{thm append_self_conv} \\ |
|
150 @{thm[source]append_self_conv[of _ "[]"]} & @{thm append_self_conv[of _ "[]"]} |
|
151 \end{tabular} |
|
152 \end{center} |
|
153 Applying one theorem to another |
|
154 by unifying the premise of one theorem with the conclusion of another: |
|
155 \begin{center} |
|
156 \begin{tabular}{@ {}ll@ {}} |
|
157 @{thm[source]sym} & @{thm sym} \\ |
|
158 @{thm[source]sym[OF append_self_conv]} & @{thm sym[OF append_self_conv]}\\ |
|
159 @{thm[source]append_self_conv[THEN sym]} & @{thm append_self_conv[THEN sym]} |
|
160 \end{tabular} |
|
161 \end{center} |
|
162 *} |
|
163 |
|
164 (*<*) |
|
165 thm append_self_conv |
|
166 thm append_self_conv[of _ "[]"] |
98 thm sym |
167 thm sym |
99 thm sym[OF rev_rev_ident] |
168 thm sym[OF append_self_conv] |
100 thm rev_rev_ident[THEN sym] |
169 thm append_self_conv[THEN sym] |
101 |
170 (*>*) |
102 |
171 |
103 subsection{*Further Useful Methods*} |
172 subsection{*Further Useful Methods*} |
104 |
173 |
105 lemma "n \<le> 1 \<and> n \<noteq> 0 \<Longrightarrow> n^n = n" |
174 lemma "n \<le> 1 \<and> n \<noteq> 0 \<Longrightarrow> n^n = n" |
106 apply(subgoal_tac "n=1") |
175 apply(subgoal_tac "n=1") |
107 apply simp |
176 apply simp |
108 apply arith |
177 apply arith |
109 done |
178 done |
110 |
179 |
|
180 text{* And a cute example: *} |
111 |
181 |
112 lemma "\<lbrakk> 2 \<in> Q; sqrt 2 \<notin> Q; |
182 lemma "\<lbrakk> 2 \<in> Q; sqrt 2 \<notin> Q; |
113 \<forall>x y z. sqrt x * sqrt x = x \<and> |
183 \<forall>x y z. sqrt x * sqrt x = x \<and> |
114 x ^ 2 = x * x \<and> |
184 x ^ 2 = x * x \<and> |
115 (x ^ y) ^ z = x ^ (y*z) |
185 (x ^ y) ^ z = x ^ (y*z) |