src/HOL/Isar_Examples/Peirce.thy
author wenzelm
Mon, 21 Jun 2010 11:24:19 +0200
changeset 37380 35815ce9218a
parent 33026 8f35633c4922
child 37671 fa53d267dab3
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33026
8f35633c4922 modernized session Isar_Examples;
wenzelm
parents: 31758
diff changeset
     1
(*  Title:      HOL/Isar_Examples/Peirce.thy
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     3
*)
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     4
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
     5
header {* Peirce's Law *}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     6
31758
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
     7
theory Peirce
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
     8
imports Main
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
     9
begin
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    10
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    11
text {*
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    12
 We consider Peirce's Law: $((A \impl B) \impl A) \impl A$.  This is
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    13
 an inherently non-intuitionistic statement, so its proof will
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    14
 certainly involve some form of classical contradiction.
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    15
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    16
 The first proof is again a well-balanced combination of plain
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    17
 backward and forward reasoning.  The actual classical step is where
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    18
 the negated goal may be introduced as additional assumption.  This
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    19
 eventually leads to a contradiction.\footnote{The rule involved there
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    20
 is negation elimination; it holds in intuitionistic logic as well.}
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    21
*}
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    22
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    23
theorem "((A --> B) --> A) --> A"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    24
proof
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    25
  assume "(A --> B) --> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    26
  show A
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    27
  proof (rule classical)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    28
    assume "~ A"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    29
    have "A --> B"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    30
    proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    31
      assume A
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    32
      with `~ A` show B by contradiction
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    33
    qed
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    34
    with `(A --> B) --> A` show A ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    35
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    36
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    37
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    38
text {*
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    39
 In the subsequent version the reasoning is rearranged by means of
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    40
 ``weak assumptions'' (as introduced by \isacommand{presume}).  Before
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    41
 assuming the negated goal $\neg A$, its intended consequence $A \impl
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    42
 B$ is put into place in order to solve the main problem.
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    43
 Nevertheless, we do not get anything for free, but have to establish
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    44
 $A \impl B$ later on.  The overall effect is that of a logical
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    45
 \emph{cut}.
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    46
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    47
 Technically speaking, whenever some goal is solved by
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    48
 \isacommand{show} in the context of weak assumptions then the latter
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    49
 give rise to new subgoals, which may be established separately.  In
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    50
 contrast, strong assumptions (as introduced by \isacommand{assume})
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    51
 are solved immediately.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    52
*}
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    53
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    54
theorem "((A --> B) --> A) --> A"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    55
proof
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    56
  assume "(A --> B) --> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    57
  show A
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    58
  proof (rule classical)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    59
    presume "A --> B"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    60
    with `(A --> B) --> A` show A ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    61
  next
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    62
    assume "~ A"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    63
    show "A --> B"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    64
    proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    65
      assume A
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    66
      with `~ A` show B by contradiction
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    67
    qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    68
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    69
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    70
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    71
text {*
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    72
 Note that the goals stemming from weak assumptions may be even left
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    73
 until qed time, where they get eventually solved ``by assumption'' as
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    74
 well.  In that case there is really no fundamental difference between
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    75
 the two kinds of assumptions, apart from the order of reducing the
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    76
 individual parts of the proof configuration.
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    77
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    78
 Nevertheless, the ``strong'' mode of plain assumptions is quite
7982
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    79
 important in practice to achieve robustness of proof text
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    80
 interpretation.  By forcing both the conclusion \emph{and} the
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7860
diff changeset
    81
 assumptions to unify with the pending goal to be solved, goal
c007f801cd59 improved presentation;
wenzelm
parents: 7860
diff changeset
    82
 selection becomes quite deterministic.  For example, decomposition
7982
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    83
 with rules of the ``case-analysis'' type usually gives rise to
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    84
 several goals that only differ in there local contexts.  With strong
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    85
 assumptions these may be still solved in any order in a predictable
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    86
 way, while weak ones would quickly lead to great confusion,
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    87
 eventually demanding even some backtracking.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    88
*}
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    89
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    90
end