| 26322 |      1 | (*  Title:      FOLP/ex/Intuitionistic.thy
 | 
|  |      2 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      3 |     Copyright   1991  University of Cambridge
 | 
|  |      4 | 
 | 
|  |      5 | Intuitionistic First-Order Logic.
 | 
|  |      6 | 
 | 
|  |      7 | Single-step commands:
 | 
|  |      8 | by (IntPr.step_tac 1)
 | 
|  |      9 | by (biresolve_tac safe_brls 1);
 | 
|  |     10 | by (biresolve_tac haz_brls 1);
 | 
|  |     11 | by (assume_tac 1);
 | 
|  |     12 | by (IntPr.safe_tac 1);
 | 
|  |     13 | by (IntPr.mp_tac 1);
 | 
|  |     14 | by (IntPr.fast_tac 1);
 | 
|  |     15 | *)
 | 
|  |     16 | 
 | 
|  |     17 | (*Note: for PROPOSITIONAL formulae...
 | 
|  |     18 |   ~A is classically provable iff it is intuitionistically provable.  
 | 
|  |     19 |   Therefore A is classically provable iff ~~A is intuitionistically provable.
 | 
|  |     20 | 
 | 
|  |     21 | Let Q be the conjuction of the propositions A|~A, one for each atom A in
 | 
|  |     22 | P.  If P is provable classically, then clearly P&Q is provable
 | 
|  |     23 | intuitionistically, so ~~(P&Q) is also provable intuitionistically.
 | 
|  |     24 | The latter is intuitionistically equivalent to ~~P&~~Q, hence to ~~P,
 | 
|  |     25 | since ~~Q is intuitionistically provable.  Finally, if P is a negation then
 | 
|  |     26 | ~~P is intuitionstically equivalent to P.  [Andy Pitts]
 | 
|  |     27 | *)
 | 
|  |     28 | 
 | 
|  |     29 | theory Intuitionistic
 | 
|  |     30 | imports IFOLP
 | 
|  |     31 | begin
 | 
|  |     32 | 
 | 
| 36319 |     33 | schematic_lemma "?p : ~~(P&Q) <-> ~~P & ~~Q"
 | 
| 26322 |     34 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     35 | 
 | 
| 36319 |     36 | schematic_lemma "?p : ~~~P <-> ~P"
 | 
| 26322 |     37 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     38 | 
 | 
| 36319 |     39 | schematic_lemma "?p : ~~((P --> Q | R)  -->  (P-->Q) | (P-->R))"
 | 
| 26322 |     40 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     41 | 
 | 
| 36319 |     42 | schematic_lemma "?p : (P<->Q) <-> (Q<->P)"
 | 
| 26322 |     43 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     44 | 
 | 
|  |     45 | 
 | 
|  |     46 | subsection {* Lemmas for the propositional double-negation translation *}
 | 
|  |     47 | 
 | 
| 36319 |     48 | schematic_lemma "?p : P --> ~~P"
 | 
| 26322 |     49 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     50 | 
 | 
| 36319 |     51 | schematic_lemma "?p : ~~(~~P --> P)"
 | 
| 26322 |     52 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     53 | 
 | 
| 36319 |     54 | schematic_lemma "?p : ~~P & ~~(P --> Q) --> ~~Q"
 | 
| 26322 |     55 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     56 | 
 | 
|  |     57 | 
 | 
|  |     58 | subsection {* The following are classically but not constructively valid *}
 | 
|  |     59 | 
 | 
|  |     60 | (*The attempt to prove them terminates quickly!*)
 | 
| 36319 |     61 | schematic_lemma "?p : ((P-->Q) --> P)  -->  P"
 | 
| 26322 |     62 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |     63 |   oops
 | 
|  |     64 | 
 | 
| 36319 |     65 | schematic_lemma "?p : (P&Q-->R)  -->  (P-->R) | (Q-->R)"
 | 
| 26322 |     66 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |     67 |   oops
 | 
|  |     68 | 
 | 
|  |     69 | 
 | 
|  |     70 | subsection {* Intuitionistic FOL: propositional problems based on Pelletier *}
 | 
|  |     71 | 
 | 
|  |     72 | text "Problem ~~1"
 | 
| 36319 |     73 | schematic_lemma "?p : ~~((P-->Q)  <->  (~Q --> ~P))"
 | 
| 26322 |     74 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     75 | 
 | 
|  |     76 | text "Problem ~~2"
 | 
| 36319 |     77 | schematic_lemma "?p : ~~(~~P  <->  P)"
 | 
| 26322 |     78 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     79 | 
 | 
|  |     80 | text "Problem 3"
 | 
| 36319 |     81 | schematic_lemma "?p : ~(P-->Q) --> (Q-->P)"
 | 
| 26322 |     82 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     83 | 
 | 
|  |     84 | text "Problem ~~4"
 | 
| 36319 |     85 | schematic_lemma "?p : ~~((~P-->Q)  <->  (~Q --> P))"
 | 
| 26322 |     86 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     87 | 
 | 
|  |     88 | text "Problem ~~5"
 | 
| 36319 |     89 | schematic_lemma "?p : ~~((P|Q-->P|R) --> P|(Q-->R))"
 | 
| 26322 |     90 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     91 | 
 | 
|  |     92 | text "Problem ~~6"
 | 
| 36319 |     93 | schematic_lemma "?p : ~~(P | ~P)"
 | 
| 26322 |     94 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     95 | 
 | 
|  |     96 | text "Problem ~~7"
 | 
| 36319 |     97 | schematic_lemma "?p : ~~(P | ~~~P)"
 | 
| 26322 |     98 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |     99 | 
 | 
|  |    100 | text "Problem ~~8.  Peirce's law"
 | 
| 36319 |    101 | schematic_lemma "?p : ~~(((P-->Q) --> P)  -->  P)"
 | 
| 26322 |    102 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    103 | 
 | 
|  |    104 | text "Problem 9"
 | 
| 36319 |    105 | schematic_lemma "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
 | 
| 26322 |    106 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    107 | 
 | 
|  |    108 | text "Problem 10"
 | 
| 36319 |    109 | schematic_lemma "?p : (Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
 | 
| 26322 |    110 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    111 | 
 | 
|  |    112 | text "11.  Proved in each direction (incorrectly, says Pelletier!!) "
 | 
| 36319 |    113 | schematic_lemma "?p : P<->P"
 | 
| 26322 |    114 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    115 | 
 | 
|  |    116 | text "Problem ~~12.  Dijkstra's law  "
 | 
| 36319 |    117 | schematic_lemma "?p : ~~(((P <-> Q) <-> R)  <->  (P <-> (Q <-> R)))"
 | 
| 26322 |    118 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    119 | 
 | 
| 36319 |    120 | schematic_lemma "?p : ((P <-> Q) <-> R)  -->  ~~(P <-> (Q <-> R))"
 | 
| 26322 |    121 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    122 | 
 | 
|  |    123 | text "Problem 13.  Distributive law"
 | 
| 36319 |    124 | schematic_lemma "?p : P | (Q & R)  <-> (P | Q) & (P | R)"
 | 
| 26322 |    125 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    126 | 
 | 
|  |    127 | text "Problem ~~14"
 | 
| 36319 |    128 | schematic_lemma "?p : ~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
 | 
| 26322 |    129 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    130 | 
 | 
|  |    131 | text "Problem ~~15"
 | 
| 36319 |    132 | schematic_lemma "?p : ~~((P --> Q) <-> (~P | Q))"
 | 
| 26322 |    133 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    134 | 
 | 
|  |    135 | text "Problem ~~16"
 | 
| 36319 |    136 | schematic_lemma "?p : ~~((P-->Q) | (Q-->P))"
 | 
| 26322 |    137 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    138 | 
 | 
|  |    139 | text "Problem ~~17"
 | 
| 36319 |    140 | schematic_lemma "?p : ~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
 | 
| 26322 |    141 |   by (tactic {* IntPr.fast_tac 1 *})  -- slow
 | 
|  |    142 | 
 | 
|  |    143 | 
 | 
|  |    144 | subsection {* Examples with quantifiers *}
 | 
|  |    145 | 
 | 
|  |    146 | text "The converse is classical in the following implications..."
 | 
|  |    147 | 
 | 
| 36319 |    148 | schematic_lemma "?p : (EX x. P(x)-->Q)  -->  (ALL x. P(x)) --> Q"
 | 
| 26322 |    149 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    150 | 
 | 
| 36319 |    151 | schematic_lemma "?p : ((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
 | 
| 26322 |    152 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    153 | 
 | 
| 36319 |    154 | schematic_lemma "?p : ((ALL x. ~P(x))-->Q)  -->  ~ (ALL x. ~ (P(x)|Q))"
 | 
| 26322 |    155 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    156 | 
 | 
| 36319 |    157 | schematic_lemma "?p : (ALL x. P(x)) | Q  -->  (ALL x. P(x) | Q)"
 | 
| 26322 |    158 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    159 | 
 | 
| 36319 |    160 | schematic_lemma "?p : (EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
 | 
| 26322 |    161 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    162 | 
 | 
|  |    163 | 
 | 
|  |    164 | text "The following are not constructively valid!"
 | 
|  |    165 | text "The attempt to prove them terminates quickly!"
 | 
|  |    166 | 
 | 
| 36319 |    167 | schematic_lemma "?p : ((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
 | 
| 26322 |    168 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |    169 |   oops
 | 
|  |    170 | 
 | 
| 36319 |    171 | schematic_lemma "?p : (P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
 | 
| 26322 |    172 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |    173 |   oops
 | 
|  |    174 | 
 | 
| 36319 |    175 | schematic_lemma "?p : (ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
 | 
| 26322 |    176 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |    177 |   oops
 | 
|  |    178 | 
 | 
| 36319 |    179 | schematic_lemma "?p : (ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
 | 
| 26322 |    180 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |    181 |   oops
 | 
|  |    182 | 
 | 
|  |    183 | (*Classically but not intuitionistically valid.  Proved by a bug in 1986!*)
 | 
| 36319 |    184 | schematic_lemma "?p : EX x. Q(x) --> (ALL x. Q(x))"
 | 
| 26322 |    185 |   apply (tactic {* IntPr.fast_tac 1 *})?
 | 
|  |    186 |   oops
 | 
|  |    187 | 
 | 
|  |    188 | 
 | 
|  |    189 | subsection "Hard examples with quantifiers"
 | 
|  |    190 | 
 | 
|  |    191 | text {*
 | 
|  |    192 |   The ones that have not been proved are not known to be valid!
 | 
|  |    193 |   Some will require quantifier duplication -- not currently available.
 | 
|  |    194 | *}
 | 
|  |    195 | 
 | 
|  |    196 | text "Problem ~~18"
 | 
| 36319 |    197 | schematic_lemma "?p : ~~(EX y. ALL x. P(y)-->P(x))" oops
 | 
| 26322 |    198 | (*NOT PROVED*)
 | 
|  |    199 | 
 | 
|  |    200 | text "Problem ~~19"
 | 
| 36319 |    201 | schematic_lemma "?p : ~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))" oops
 | 
| 26322 |    202 | (*NOT PROVED*)
 | 
|  |    203 | 
 | 
|  |    204 | text "Problem 20"
 | 
| 36319 |    205 | schematic_lemma "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))      
 | 
| 26322 |    206 |     --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
 | 
|  |    207 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    208 | 
 | 
|  |    209 | text "Problem 21"
 | 
| 36319 |    210 | schematic_lemma "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))" oops
 | 
| 26322 |    211 | (*NOT PROVED*)
 | 
|  |    212 | 
 | 
|  |    213 | text "Problem 22"
 | 
| 36319 |    214 | schematic_lemma "?p : (ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
 | 
| 26322 |    215 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    216 | 
 | 
|  |    217 | text "Problem ~~23"
 | 
| 36319 |    218 | schematic_lemma "?p : ~~ ((ALL x. P | Q(x))  <->  (P | (ALL x. Q(x))))"
 | 
| 26322 |    219 |   by (tactic {* IntPr.fast_tac 1 *})
 | 
|  |    220 | 
 | 
|  |    221 | text "Problem 24"
 | 
| 36319 |    222 | schematic_lemma "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &   
 | 
| 26322 |    223 |      (~(EX x. P(x)) --> (EX x. Q(x))) & (ALL x. Q(x)|R(x) --> S(x))   
 | 
|  |    224 |     --> ~~(EX x. P(x)&R(x))"
 | 
|  |    225 | (*Not clear why fast_tac, best_tac, ASTAR and ITER_DEEPEN all take forever*)
 | 
|  |    226 |   apply (tactic "IntPr.safe_tac")
 | 
|  |    227 |   apply (erule impE)
 | 
|  |    228 |    apply (tactic "IntPr.fast_tac 1")
 | 
|  |    229 |   apply (tactic "IntPr.fast_tac 1")
 | 
|  |    230 |   done
 | 
|  |    231 | 
 | 
|  |    232 | text "Problem 25"
 | 
| 36319 |    233 | schematic_lemma "?p : (EX x. P(x)) &   
 | 
| 26322 |    234 |         (ALL x. L(x) --> ~ (M(x) & R(x))) &   
 | 
|  |    235 |         (ALL x. P(x) --> (M(x) & L(x))) &    
 | 
|  |    236 |         ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))   
 | 
|  |    237 |     --> (EX x. Q(x)&P(x))"
 | 
|  |    238 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    239 | 
 | 
|  |    240 | text "Problem 29.  Essentially the same as Principia Mathematica *11.71"
 | 
| 36319 |    241 | schematic_lemma "?p : (EX x. P(x)) & (EX y. Q(y))   
 | 
| 26322 |    242 |     --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y))   <->      
 | 
|  |    243 |          (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"
 | 
|  |    244 |   by (tactic "IntPr.fast_tac 1")
 | 
|  |    245 | 
 | 
|  |    246 | text "Problem ~~30"
 | 
| 36319 |    247 | schematic_lemma "?p : (ALL x. (P(x) | Q(x)) --> ~ R(x)) &  
 | 
| 26322 |    248 |         (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))   
 | 
|  |    249 |     --> (ALL x. ~~S(x))"
 | 
|  |    250 |   by (tactic "IntPr.fast_tac 1")
 | 
|  |    251 | 
 | 
|  |    252 | text "Problem 31"
 | 
| 36319 |    253 | schematic_lemma "?p : ~(EX x. P(x) & (Q(x) | R(x))) &  
 | 
| 26322 |    254 |         (EX x. L(x) & P(x)) &  
 | 
|  |    255 |         (ALL x. ~ R(x) --> M(x))   
 | 
|  |    256 |     --> (EX x. L(x) & M(x))"
 | 
|  |    257 |   by (tactic "IntPr.fast_tac 1")
 | 
|  |    258 | 
 | 
|  |    259 | text "Problem 32"
 | 
| 36319 |    260 | schematic_lemma "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) &  
 | 
| 26322 |    261 |         (ALL x. S(x) & R(x) --> L(x)) &  
 | 
|  |    262 |         (ALL x. M(x) --> R(x))   
 | 
|  |    263 |     --> (ALL x. P(x) & M(x) --> L(x))"
 | 
|  |    264 |   by (tactic "IntPr.best_tac 1") -- slow
 | 
|  |    265 | 
 | 
|  |    266 | text "Problem 39"
 | 
| 36319 |    267 | schematic_lemma "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
 | 
| 26322 |    268 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    269 | 
 | 
|  |    270 | text "Problem 40.  AMENDED"
 | 
| 36319 |    271 | schematic_lemma "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) -->   
 | 
| 26322 |    272 |               ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"
 | 
|  |    273 |   by (tactic "IntPr.best_tac 1") -- slow
 | 
|  |    274 | 
 | 
|  |    275 | text "Problem 44"
 | 
| 36319 |    276 | schematic_lemma "?p : (ALL x. f(x) -->                                    
 | 
| 26322 |    277 |               (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y))))  &        
 | 
|  |    278 |               (EX x. j(x) & (ALL y. g(y) --> h(x,y)))                    
 | 
|  |    279 |               --> (EX x. j(x) & ~f(x))"
 | 
|  |    280 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    281 | 
 | 
|  |    282 | text "Problem 48"
 | 
| 36319 |    283 | schematic_lemma "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"
 | 
| 26322 |    284 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    285 | 
 | 
|  |    286 | text "Problem 51"
 | 
| 36319 |    287 | schematic_lemma
 | 
| 26322 |    288 |     "?p : (EX z w. ALL x y. P(x,y) <->  (x=z & y=w)) -->   
 | 
|  |    289 |      (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"
 | 
|  |    290 |   by (tactic "IntPr.best_tac 1") -- {*60 seconds*}
 | 
|  |    291 | 
 | 
|  |    292 | text "Problem 56"
 | 
| 36319 |    293 | schematic_lemma "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
 | 
| 26322 |    294 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    295 | 
 | 
|  |    296 | text "Problem 57"
 | 
| 36319 |    297 | schematic_lemma
 | 
| 26322 |    298 |     "?p : P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &  
 | 
|  |    299 |      (ALL x y z. P(x,y) & P(y,z) --> P(x,z))    -->   P(f(a,b), f(a,c))"
 | 
|  |    300 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    301 | 
 | 
|  |    302 | text "Problem 60"
 | 
| 36319 |    303 | schematic_lemma "?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
 | 
| 26322 |    304 |   by (tactic "IntPr.best_tac 1")
 | 
|  |    305 | 
 | 
|  |    306 | end
 |