src/FOL/ex/Intuitionistic.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 14239 af2a9e68bea9
child 31974 e81979a703a4
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     1
(*  Title:      FOL/ex/Intuitionistic
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     2
    ID:         $Id$
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     5
*)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     6
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     7
header{*Intuitionistic First-Order Logic*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14239
diff changeset
     9
theory Intuitionistic imports IFOL begin
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    10
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    11
(*
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    12
Single-step ML commands:
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    13
by (IntPr.step_tac 1)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    14
by (biresolve_tac safe_brls 1);
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    15
by (biresolve_tac haz_brls 1);
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    16
by (assume_tac 1);
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    17
by (IntPr.safe_tac 1);
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    18
by (IntPr.mp_tac 1);
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    19
by (IntPr.fast_tac 1);
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    20
*)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    21
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    22
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    23
text{*Metatheorem (for \emph{propositional} formulae):
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    24
  $P$ is classically provable iff $\neg\neg P$ is intuitionistically provable.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    25
  Therefore $\neg P$ is classically provable iff it is intuitionistically
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    26
  provable.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    27
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    28
Proof: Let $Q$ be the conjuction of the propositions $A\vee\neg A$, one for
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    29
each atom $A$ in $P$.  Now $\neg\neg Q$ is intuitionistically provable because
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    30
$\neg\neg(A\vee\neg A)$ is and because double-negation distributes over
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    31
conjunction.  If $P$ is provable classically, then clearly $Q\rightarrow P$ is
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    32
provable intuitionistically, so $\neg\neg(Q\rightarrow P)$ is also provable
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    33
intuitionistically.  The latter is intuitionistically equivalent to $\neg\neg
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    34
Q\rightarrow\neg\neg P$, hence to $\neg\neg P$, since $\neg\neg Q$ is
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    35
intuitionistically provable.  Finally, if $P$ is a negation then $\neg\neg P$
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    36
is intuitionstically equivalent to $P$.  [Andy Pitts] *}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    37
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    38
lemma "~~(P&Q) <-> ~~P & ~~Q"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    39
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    40
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    41
lemma "~~ ((~P --> Q) --> (~P --> ~Q) --> P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    42
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    43
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    44
text{*Double-negation does NOT distribute over disjunction*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    45
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    46
lemma "~~(P-->Q)  <-> (~~P --> ~~Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    47
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    48
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    49
lemma "~~~P <-> ~P"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    50
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    51
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    52
lemma "~~((P --> Q | R)  -->  (P-->Q) | (P-->R))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    53
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    54
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    55
lemma "(P<->Q) <-> (Q<->P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    56
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    57
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    58
lemma "((P --> (Q | (Q-->R))) --> R) --> R"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    59
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    60
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    61
lemma "(((G-->A) --> J) --> D --> E) --> (((H-->B)-->I)-->C-->J)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    62
      --> (A-->H) --> F --> G --> (((C-->B)-->I)-->D)-->(A-->C)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    63
      --> (((F-->A)-->B) --> I) --> E"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    64
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    65
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    66
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    67
text{*Lemmas for the propositional double-negation translation*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    68
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    69
lemma "P --> ~~P"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    70
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    71
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    72
lemma "~~(~~P --> P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    73
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    74
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    75
lemma "~~P & ~~(P --> Q) --> ~~Q"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    76
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    77
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    78
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    79
text{*The following are classically but not constructively valid.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    80
      The attempt to prove them terminates quickly!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    81
lemma "((P-->Q) --> P)  -->  P"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    82
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    83
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    84
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    85
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    86
lemma "(P&Q-->R)  -->  (P-->R) | (Q-->R)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    87
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    88
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    89
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    90
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    91
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    92
subsection{*de Bruijn formulae*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    93
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    94
text{*de Bruijn formula with three predicates*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    95
lemma "((P<->Q) --> P&Q&R) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    96
               ((Q<->R) --> P&Q&R) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    97
               ((R<->P) --> P&Q&R) --> P&Q&R"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    98
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    99
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   100
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   101
text{*de Bruijn formula with five predicates*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   102
lemma "((P<->Q) --> P&Q&R&S&T) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   103
               ((Q<->R) --> P&Q&R&S&T) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   104
               ((R<->S) --> P&Q&R&S&T) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   105
               ((S<->T) --> P&Q&R&S&T) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   106
               ((T<->P) --> P&Q&R&S&T) --> P&Q&R&S&T"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   107
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   108
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   109
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   110
(*** Problems from of Sahlin, Franzen and Haridi,
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   111
     An Intuitionistic Predicate Logic Theorem Prover.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   112
     J. Logic and Comp. 2 (5), October 1992, 619-656.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   113
***)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   114
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   115
text{*Problem 1.1*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   116
lemma "(ALL x. EX y. ALL z. p(x) & q(y) & r(z)) <->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   117
      (ALL z. EX y. ALL x. p(x) & q(y) & r(z))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   118
by (tactic{*IntPr.best_dup_tac 1*})  --{*SLOW*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   119
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   120
text{*Problem 3.1*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   121
lemma "~ (EX x. ALL y. mem(y,x) <-> ~ mem(x,x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   122
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   123
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   124
text{*Problem 4.1: hopeless!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   125
lemma "(ALL x. p(x) --> p(h(x)) | p(g(x))) & (EX x. p(x)) & (ALL x. ~p(h(x)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   126
      --> (EX x. p(g(g(g(g(g(x)))))))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   127
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   128
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   129
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   130
subsection{*Intuitionistic FOL: propositional problems based on Pelletier.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   131
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   132
text{*~~1*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   133
lemma "~~((P-->Q)  <->  (~Q --> ~P))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   134
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   135
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   136
text{*~~2*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   137
lemma "~~(~~P  <->  P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   138
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   139
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   140
text{*3*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   141
lemma "~(P-->Q) --> (Q-->P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   142
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   143
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   144
text{*~~4*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   145
lemma "~~((~P-->Q)  <->  (~Q --> P))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   146
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   147
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   148
text{*~~5*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   149
lemma "~~((P|Q-->P|R) --> P|(Q-->R))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   150
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   151
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   152
text{*~~6*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   153
lemma "~~(P | ~P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   154
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   155
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   156
text{*~~7*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   157
lemma "~~(P | ~~~P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   158
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   159
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   160
text{*~~8.  Peirce's law*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   161
lemma "~~(((P-->Q) --> P)  -->  P)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   162
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   163
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   164
text{*9*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   165
lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   166
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   167
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   168
text{*10*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   169
lemma "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   170
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   171
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   172
subsection{*11.  Proved in each direction (incorrectly, says Pelletier!!) *}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   173
lemma "P<->P"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   174
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   175
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   176
text{*~~12.  Dijkstra's law  *}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   177
lemma "~~(((P <-> Q) <-> R)  <->  (P <-> (Q <-> R)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   178
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   179
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   180
lemma "((P <-> Q) <-> R)  -->  ~~(P <-> (Q <-> R))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   181
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   182
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   183
text{*13.  Distributive law*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   184
lemma "P | (Q & R)  <-> (P | Q) & (P | R)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   185
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   186
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   187
text{*~~14*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   188
lemma "~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   189
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   190
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   191
text{*~~15*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   192
lemma "~~((P --> Q) <-> (~P | Q))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   193
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   194
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   195
text{*~~16*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   196
lemma "~~((P-->Q) | (Q-->P))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   197
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   198
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   199
text{*~~17*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   200
lemma "~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   201
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   202
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   203
text{*Dijkstra's "Golden Rule"*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   204
lemma "(P&Q) <-> P <-> Q <-> (P|Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   205
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   206
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   207
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   208
subsection{*****Examples with quantifiers*****}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   209
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   210
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   211
subsection{*The converse is classical in the following implications...*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   212
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   213
lemma "(EX x. P(x)-->Q)  -->  (ALL x. P(x)) --> Q"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   214
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   215
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   216
lemma "((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   217
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   218
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   219
lemma "((ALL x. ~P(x))-->Q)  -->  ~ (ALL x. ~ (P(x)|Q))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   220
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   221
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   222
lemma "(ALL x. P(x)) | Q  -->  (ALL x. P(x) | Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   223
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   224
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   225
lemma "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   226
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   227
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   228
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   229
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   230
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   231
subsection{*The following are not constructively valid!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   232
text{*The attempt to prove them terminates quickly!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   233
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   234
lemma "((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   235
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   236
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   237
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   238
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   239
lemma "(P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   240
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   241
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   242
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   243
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   244
lemma "(ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   245
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   246
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   247
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   248
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   249
lemma "(ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   250
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   251
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   252
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   253
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   254
text{*Classically but not intuitionistically valid.  Proved by a bug in 1986!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   255
lemma "EX x. Q(x) --> (ALL x. Q(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   256
apply (tactic{*IntPr.fast_tac 1*} | -)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   257
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   258
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   259
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   260
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   261
subsection{*Hard examples with quantifiers*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   262
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   263
text{*The ones that have not been proved are not known to be valid!
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   264
  Some will require quantifier duplication -- not currently available*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   265
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   266
text{*~~18*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   267
lemma "~~(EX y. ALL x. P(y)-->P(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   268
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   269
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   270
text{*~~19*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   271
lemma "~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   272
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   273
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   274
text{*20*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   275
lemma "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   276
    --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   277
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   278
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   279
text{*21*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   280
lemma "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   281
oops --{*NOT PROVED; needs quantifier duplication*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   282
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   283
text{*22*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   284
lemma "(ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   285
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   286
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   287
text{*~~23*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   288
lemma "~~ ((ALL x. P | Q(x))  <->  (P | (ALL x. Q(x))))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   289
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   290
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   291
text{*24*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   292
lemma "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   293
     (~(EX x. P(x)) --> (EX x. Q(x))) & (ALL x. Q(x)|R(x) --> S(x))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   294
    --> ~~(EX x. P(x)&R(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   295
txt{*Not clear why @{text fast_tac}, @{text best_tac}, @{text ASTAR} and 
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   296
    @{text ITER_DEEPEN} all take forever*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   297
apply (tactic{* IntPr.safe_tac*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   298
apply (erule impE)
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   299
apply (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   300
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   301
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   302
text{*25*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   303
lemma "(EX x. P(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   304
        (ALL x. L(x) --> ~ (M(x) & R(x))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   305
        (ALL x. P(x) --> (M(x) & L(x))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   306
        ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   307
    --> (EX x. Q(x)&P(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   308
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   309
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   310
text{*~~26*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   311
lemma "(~~(EX x. p(x)) <-> ~~(EX x. q(x))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   312
      (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   313
  --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   314
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   315
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   316
text{*27*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   317
lemma "(EX x. P(x) & ~Q(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   318
              (ALL x. P(x) --> R(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   319
              (ALL x. M(x) & L(x) --> P(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   320
              ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   321
          --> (ALL x. M(x) --> ~L(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   322
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   323
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   324
text{*~~28.  AMENDED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   325
lemma "(ALL x. P(x) --> (ALL x. Q(x))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   326
        (~~(ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   327
        (~~(EX x. S(x)) --> (ALL x. L(x) --> M(x)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   328
    --> (ALL x. P(x) & L(x) --> M(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   329
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   330
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   331
text{*29.  Essentially the same as Principia Mathematica *11.71*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   332
lemma "(EX x. P(x)) & (EX y. Q(y))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   333
    --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y))   <->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   334
         (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   335
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   336
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   337
text{*~~30*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   338
lemma "(ALL x. (P(x) | Q(x)) --> ~ R(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   339
        (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   340
    --> (ALL x. ~~S(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   341
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   342
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   343
text{*31*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   344
lemma "~(EX x. P(x) & (Q(x) | R(x))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   345
        (EX x. L(x) & P(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   346
        (ALL x. ~ R(x) --> M(x))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   347
    --> (EX x. L(x) & M(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   348
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   349
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   350
text{*32*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   351
lemma "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   352
        (ALL x. S(x) & R(x) --> L(x)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   353
        (ALL x. M(x) --> R(x))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   354
    --> (ALL x. P(x) & M(x) --> L(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   355
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   356
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   357
text{*~~33*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   358
lemma "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c)))  <->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   359
      (ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   360
apply (tactic{*IntPr.best_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   361
done
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   362
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   363
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   364
text{*36*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   365
lemma "(ALL x. EX y. J(x,y)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   366
      (ALL x. EX y. G(x,y)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   367
      (ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   368
  --> (ALL x. EX y. H(x,y))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   369
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   370
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   371
text{*37*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   372
lemma "(ALL z. EX w. ALL x. EX y.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   373
           ~~(P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u. Q(u,w)))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   374
        (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   375
        (~~(EX x y. Q(x,y)) --> (ALL x. R(x,x)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   376
    --> ~~(ALL x. EX y. R(x,y))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   377
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   378
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   379
text{*39*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   380
lemma "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   381
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   382
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   383
text{*40.  AMENDED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   384
lemma "(EX y. ALL x. F(x,y) <-> F(x,x)) -->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   385
              ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   386
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   387
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   388
text{*44*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   389
lemma "(ALL x. f(x) -->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   390
              (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y))))  &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   391
              (EX x. j(x) & (ALL y. g(y) --> h(x,y)))
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   392
              --> (EX x. j(x) & ~f(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   393
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   394
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   395
text{*48*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   396
lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   397
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   398
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   399
text{*51*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   400
lemma "(EX z w. ALL x y. P(x,y) <->  (x=z & y=w)) -->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   401
     (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   402
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   403
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   404
text{*52*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   405
text{*Almost the same as 51. *}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   406
lemma "(EX z w. ALL x y. P(x,y) <->  (x=z & y=w)) -->
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   407
     (EX w. ALL y. EX z. (ALL x. P(x,y) <-> x=z) <-> y=w)"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   408
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   409
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   410
text{*56*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   411
lemma "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   412
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   413
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   414
text{*57*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   415
lemma "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   416
     (ALL x y z. P(x,y) & P(y,z) --> P(x,z))    -->   P(f(a,b), f(a,c))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   417
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   418
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   419
text{*60*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   420
lemma "ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   421
by (tactic{*IntPr.fast_tac 1*})
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   422
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   423
end
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   424
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   425