src/FOL/ex/Intuitionistic.thy
author wenzelm
Sun, 02 Nov 2014 18:21:45 +0100
changeset 58889 5b7a9633cfa8
parent 51798 ad3a241def73
child 60770 240563fbf41d
permissions -rw-r--r--
modernized header uniformly as section;
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
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 51798
diff changeset
     6
section {* Intuitionistic First-Order Logic *}
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    24
text{*Metatheorem (for \emph{propositional} formulae):
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$
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    37
is intuitionstically equivalent to $P$.  [Andy Pitts] *}
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    40
by (tactic{*IntPr.fast_tac @{context} 1*})
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    43
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    44
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    45
text{*Double-negation does NOT distribute over disjunction*}
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    48
by (tactic{*IntPr.fast_tac @{context} 1*})
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    51
by (tactic{*IntPr.fast_tac @{context} 1*})
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    54
by (tactic{*IntPr.fast_tac @{context} 1*})
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    57
by (tactic{*IntPr.fast_tac @{context} 1*})
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    60
by (tactic{*IntPr.fast_tac @{context} 1*})
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    65
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    66
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    67
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    68
text{*Lemmas for the propositional double-negation translation*}
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    71
by (tactic{*IntPr.fast_tac @{context} 1*})
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    74
by (tactic{*IntPr.fast_tac @{context} 1*})
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    77
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    78
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    79
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    80
text{*The following are classically but not constructively valid.
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    81
      The attempt to prove them terminates quickly!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    82
lemma "((P-->Q) --> P)  -->  P"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    83
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    84
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    88
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    89
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    93
subsection{*de Bruijn formulae*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    94
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
    95
text{*de Bruijn formula with three predicates*}
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
    99
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   100
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   101
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   102
text{*de Bruijn formula with five predicates*}
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   108
by (tactic{*IntPr.fast_tac @{context} 1*})
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   116
text{*Problem 1.1*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   119
by (tactic{*IntPr.best_dup_tac @{context} 1*})  --{*SLOW*}
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   120
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   121
text{*Problem 3.1*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   122
lemma "~ (EX x. ALL y. mem(y,x) <-> ~ mem(x,x))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   123
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   124
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   125
text{*Problem 4.1: hopeless!*}
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   131
subsection{*Intuitionistic FOL: propositional problems based on Pelletier.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   132
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   133
text{*~~1*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   134
lemma "~~((P-->Q)  <->  (~Q --> ~P))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   135
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   136
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   137
text{*~~2*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   138
lemma "~~(~~P  <->  P)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   139
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   140
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   141
text{*3*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   142
lemma "~(P-->Q) --> (Q-->P)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   143
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   144
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   145
text{*~~4*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   146
lemma "~~((~P-->Q)  <->  (~Q --> P))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   147
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   148
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   149
text{*~~5*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   150
lemma "~~((P|Q-->P|R) --> P|(Q-->R))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   151
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   152
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   153
text{*~~6*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   154
lemma "~~(P | ~P)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   155
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   156
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   157
text{*~~7*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   158
lemma "~~(P | ~~~P)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   159
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   160
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   161
text{*~~8.  Peirce's law*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   162
lemma "~~(((P-->Q) --> P)  -->  P)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   163
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   164
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   165
text{*9*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   166
lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   167
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   168
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   169
text{*10*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   170
lemma "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   171
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   172
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   173
subsection{*11.  Proved in each direction (incorrectly, says Pelletier!!) *}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   174
lemma "P<->P"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   175
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   176
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   177
text{*~~12.  Dijkstra's law  *}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   178
lemma "~~(((P <-> Q) <-> R)  <->  (P <-> (Q <-> R)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   179
by (tactic{*IntPr.fast_tac @{context} 1*})
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   182
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   183
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   184
text{*13.  Distributive law*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   185
lemma "P | (Q & R)  <-> (P | Q) & (P | R)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   186
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   187
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   188
text{*~~14*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   189
lemma "~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   190
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   191
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   192
text{*~~15*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   193
lemma "~~((P --> Q) <-> (~P | Q))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   194
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   195
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   196
text{*~~16*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   197
lemma "~~((P-->Q) | (Q-->P))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   198
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   199
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   200
text{*~~17*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   201
lemma "~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   202
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   203
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   204
text{*Dijkstra's "Golden Rule"*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   205
lemma "(P&Q) <-> P <-> Q <-> (P|Q)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   206
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   207
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   208
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   209
subsection{*****Examples with quantifiers*****}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   210
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   211
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   212
subsection{*The converse is classical in the following implications...*}
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"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   215
by (tactic{*IntPr.fast_tac @{context} 1*})
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   218
by (tactic{*IntPr.fast_tac @{context} 1*})
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   221
by (tactic{*IntPr.fast_tac @{context} 1*})
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   224
by (tactic{*IntPr.fast_tac @{context} 1*})
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)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   227
by (tactic{*IntPr.fast_tac @{context} 1*})
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   232
subsection{*The following are not constructively valid!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   233
text{*The attempt to prove them terminates quickly!*}
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   236
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   237
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   241
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   242
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   246
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   247
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   251
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   252
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   253
oops
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   254
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   255
text{*Classically but not intuitionistically valid.  Proved by a bug in 1986!*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   256
lemma "EX x. Q(x) --> (ALL x. Q(x))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   257
apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   258
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   262
subsection{*Hard examples with quantifiers*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   263
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   264
text{*The ones that have not been proved are not known to be valid!
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   265
  Some will require quantifier duplication -- not currently available*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   266
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   267
text{*~~18*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   268
lemma "~~(EX y. ALL x. P(y)-->P(x))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   269
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   270
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   271
text{*~~19*}
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)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   273
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   274
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   275
text{*20*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   278
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   279
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   280
text{*21*}
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))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   282
oops --{*NOT PROVED; needs quantifier duplication*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   283
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   284
text{*22*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   285
lemma "(ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   286
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   287
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   288
text{*~~23*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   289
lemma "~~ ((ALL x. P | Q(x))  <->  (P | (ALL x. Q(x))))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   290
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   291
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   292
text{*24*}
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))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   296
txt{*Not clear why @{text fast_tac}, @{text best_tac}, @{text ASTAR} and 
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   297
    @{text ITER_DEEPEN} all take forever*}
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   298
apply (tactic{* IntPr.safe_tac @{context}*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   299
apply (erule impE)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   300
apply (tactic{*IntPr.fast_tac @{context} 1*})
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   301
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   302
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   303
text{*25*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   309
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   310
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   311
text{*~~26*}
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)))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   315
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   316
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   317
text{*27*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   323
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   324
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   325
text{*~~28.  AMENDED*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   330
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   331
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   332
text{*29.  Essentially the same as Principia Mathematica *11.71*}
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)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   336
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   337
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   338
text{*~~30*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   342
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   343
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   344
text{*31*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   349
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   350
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   351
text{*32*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   356
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   357
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   358
text{*~~33*}
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))))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   361
apply (tactic{*IntPr.best_tac @{context} 1*})
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
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   365
text{*36*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   370
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   371
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   372
text{*37*}
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))"
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   378
oops  --{*NOT PROVED*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   379
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   380
text{*39*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   381
lemma "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   382
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   383
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   384
text{*40.  AMENDED*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   387
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   388
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   389
text{*44*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   394
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   395
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   396
text{*48*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   397
lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   398
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   399
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   400
text{*51*}
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   403
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   404
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   405
text{*52*}
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   406
text{*Almost the same as 51. *}
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)"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   409
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   410
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   411
text{*56*}
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)))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   413
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   414
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   415
text{*57*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   418
by (tactic{*IntPr.fast_tac @{context} 1*})
14239
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   419
af2a9e68bea9 partial conversion to Isar scripts
paulson
parents:
diff changeset
   420
text{*60*}
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))"
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 31974
diff changeset
   422
by (tactic{*IntPr.fast_tac @{context} 1*})
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