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