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