| 41959 |      1 | (*  Title:      Sequents/LK/Propositional.thy
 | 
| 21426 |      2 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      3 |     Copyright   1992  University of Cambridge
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
| 60770 |      6 | section \<open>Classical sequent calculus: examples with propositional connectives\<close>
 | 
| 21426 |      7 | 
 | 
|  |      8 | theory Propositional
 | 
| 55229 |      9 | imports "../LK"
 | 
| 21426 |     10 | begin
 | 
|  |     11 | 
 | 
| 61385 |     12 | text "absorptive laws of \<and> and \<or>"
 | 
| 21426 |     13 | 
 | 
| 61386 |     14 | lemma "\<turnstile> P \<and> P \<longleftrightarrow> P"
 | 
| 21426 |     15 |   by fast_prop
 | 
|  |     16 | 
 | 
| 61386 |     17 | lemma "\<turnstile> P \<or> P \<longleftrightarrow> P"
 | 
| 21426 |     18 |   by fast_prop
 | 
|  |     19 | 
 | 
|  |     20 | 
 | 
| 61385 |     21 | text "commutative laws of \<and> and \<or>"
 | 
| 21426 |     22 | 
 | 
| 61386 |     23 | lemma "\<turnstile> P \<and> Q \<longleftrightarrow> Q \<and> P"
 | 
| 21426 |     24 |   by fast_prop
 | 
|  |     25 | 
 | 
| 61386 |     26 | lemma "\<turnstile> P \<or> Q \<longleftrightarrow> Q \<or> P"
 | 
| 21426 |     27 |   by fast_prop
 | 
|  |     28 | 
 | 
|  |     29 | 
 | 
| 61385 |     30 | text "associative laws of \<and> and \<or>"
 | 
| 21426 |     31 | 
 | 
| 61386 |     32 | lemma "\<turnstile> (P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)"
 | 
| 21426 |     33 |   by fast_prop
 | 
|  |     34 | 
 | 
| 61386 |     35 | lemma "\<turnstile> (P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)"
 | 
| 21426 |     36 |   by fast_prop
 | 
|  |     37 | 
 | 
|  |     38 | 
 | 
| 61385 |     39 | text "distributive laws of \<and> and \<or>"
 | 
| 21426 |     40 | 
 | 
| 61386 |     41 | lemma "\<turnstile> (P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)"
 | 
| 21426 |     42 |   by fast_prop
 | 
|  |     43 | 
 | 
| 61386 |     44 | lemma "\<turnstile> (P \<or> Q) \<and> R \<longleftrightarrow> (P \<and> R) \<or> (Q \<and> R)"
 | 
| 21426 |     45 |   by fast_prop
 | 
|  |     46 | 
 | 
|  |     47 | 
 | 
|  |     48 | text "Laws involving implication"
 | 
|  |     49 | 
 | 
| 61386 |     50 | lemma "\<turnstile> (P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)"
 | 
| 21426 |     51 |   by fast_prop
 | 
|  |     52 | 
 | 
| 61386 |     53 | lemma "\<turnstile> (P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))"
 | 
| 21426 |     54 |   by fast_prop
 | 
|  |     55 | 
 | 
| 61386 |     56 | lemma "\<turnstile> (P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)"
 | 
| 21426 |     57 |   by fast_prop
 | 
|  |     58 | 
 | 
|  |     59 | 
 | 
|  |     60 | text "Classical theorems"
 | 
|  |     61 | 
 | 
| 61386 |     62 | lemma "\<turnstile> P \<or> Q \<longrightarrow> P \<or> \<not> P \<and> Q"
 | 
| 21426 |     63 |   by fast_prop
 | 
|  |     64 | 
 | 
| 61386 |     65 | lemma "\<turnstile> (P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<or> R)"
 | 
| 21426 |     66 |   by fast_prop
 | 
|  |     67 | 
 | 
| 61386 |     68 | lemma "\<turnstile> P \<and> Q \<or> \<not> P \<and> R \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> R)"
 | 
| 21426 |     69 |   by fast_prop
 | 
|  |     70 | 
 | 
| 61386 |     71 | lemma "\<turnstile> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)"
 | 
| 21426 |     72 |   by fast_prop
 | 
|  |     73 | 
 | 
|  |     74 | 
 | 
|  |     75 | (*If and only if*)
 | 
|  |     76 | 
 | 
| 61386 |     77 | lemma "\<turnstile> (P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)"
 | 
| 21426 |     78 |   by fast_prop
 | 
|  |     79 | 
 | 
| 61386 |     80 | lemma "\<turnstile> \<not> (P \<longleftrightarrow> \<not> P)"
 | 
| 21426 |     81 |   by fast_prop
 | 
|  |     82 | 
 | 
|  |     83 | 
 | 
|  |     84 | (*Sample problems from 
 | 
|  |     85 |   F. J. Pelletier, 
 | 
|  |     86 |   Seventy-Five Problems for Testing Automatic Theorem Provers,
 | 
|  |     87 |   J. Automated Reasoning 2 (1986), 191-216.
 | 
|  |     88 |   Errata, JAR 4 (1988), 236-236.
 | 
|  |     89 | *)
 | 
|  |     90 | 
 | 
|  |     91 | (*1*)
 | 
| 61386 |     92 | lemma "\<turnstile> (P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> \<not> P)"
 | 
| 21426 |     93 |   by fast_prop
 | 
|  |     94 | 
 | 
|  |     95 | (*2*)
 | 
| 61386 |     96 | lemma "\<turnstile> \<not> \<not> P \<longleftrightarrow> P"
 | 
| 21426 |     97 |   by fast_prop
 | 
|  |     98 | 
 | 
|  |     99 | (*3*)
 | 
| 61386 |    100 | lemma "\<turnstile> \<not> (P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P)"
 | 
| 21426 |    101 |   by fast_prop
 | 
|  |    102 | 
 | 
|  |    103 | (*4*)
 | 
| 61386 |    104 | lemma "\<turnstile> (\<not> P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> P)"
 | 
| 21426 |    105 |   by fast_prop
 | 
|  |    106 | 
 | 
|  |    107 | (*5*)
 | 
| 61386 |    108 | lemma "\<turnstile> ((P \<or> Q) \<longrightarrow> (P \<or> R)) \<longrightarrow> (P \<or> (Q \<longrightarrow> R))"
 | 
| 21426 |    109 |   by fast_prop
 | 
|  |    110 | 
 | 
|  |    111 | (*6*)
 | 
| 61386 |    112 | lemma "\<turnstile> P \<or> \<not> P"
 | 
| 21426 |    113 |   by fast_prop
 | 
|  |    114 | 
 | 
|  |    115 | (*7*)
 | 
| 61386 |    116 | lemma "\<turnstile> P \<or> \<not> \<not> \<not> P"
 | 
| 21426 |    117 |   by fast_prop
 | 
|  |    118 | 
 | 
|  |    119 | (*8.  Peirce's law*)
 | 
| 61386 |    120 | lemma "\<turnstile> ((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
 | 
| 21426 |    121 |   by fast_prop
 | 
|  |    122 | 
 | 
|  |    123 | (*9*)
 | 
| 61386 |    124 | lemma "\<turnstile> ((P \<or> Q) \<and> (\<not> P \<or> Q) \<and> (P \<or> \<not> Q)) \<longrightarrow> \<not> (\<not> P \<or> \<not> Q)"
 | 
| 21426 |    125 |   by fast_prop
 | 
|  |    126 | 
 | 
|  |    127 | (*10*)
 | 
| 61386 |    128 | lemma "Q \<longrightarrow> R, R \<longrightarrow> P \<and> Q, P \<longrightarrow> (Q \<or> R) \<turnstile> P \<longleftrightarrow> Q"
 | 
| 21426 |    129 |   by fast_prop
 | 
|  |    130 | 
 | 
|  |    131 | (*11.  Proved in each direction (incorrectly, says Pelletier!!)  *)
 | 
| 61386 |    132 | lemma "\<turnstile> P \<longleftrightarrow> P"
 | 
| 21426 |    133 |   by fast_prop
 | 
|  |    134 | 
 | 
|  |    135 | (*12.  "Dijkstra's law"*)
 | 
| 61386 |    136 | lemma "\<turnstile> ((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))"
 | 
| 21426 |    137 |   by fast_prop
 | 
|  |    138 | 
 | 
|  |    139 | (*13.  Distributive law*)
 | 
| 61386 |    140 | lemma "\<turnstile> P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)"
 | 
| 21426 |    141 |   by fast_prop
 | 
|  |    142 | 
 | 
|  |    143 | (*14*)
 | 
| 61386 |    144 | lemma "\<turnstile> (P \<longleftrightarrow> Q) \<longleftrightarrow> ((Q \<or> \<not> P) \<and> (\<not> Q \<or> P))"
 | 
| 21426 |    145 |   by fast_prop
 | 
|  |    146 | 
 | 
|  |    147 | (*15*)
 | 
| 61386 |    148 | lemma "\<turnstile> (P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"
 | 
| 21426 |    149 |   by fast_prop
 | 
|  |    150 | 
 | 
|  |    151 | (*16*)
 | 
| 61386 |    152 | lemma "\<turnstile> (P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"
 | 
| 21426 |    153 |   by fast_prop
 | 
|  |    154 | 
 | 
|  |    155 | (*17*)
 | 
| 61386 |    156 | lemma "\<turnstile> ((P \<and> (Q \<longrightarrow> R)) \<longrightarrow> S) \<longleftrightarrow> ((\<not> P \<or> Q \<or> S) \<and> (\<not> P \<or> \<not> R \<or> S))"
 | 
| 21426 |    157 |   by fast_prop
 | 
|  |    158 | 
 | 
|  |    159 | end
 |