src/HOL/Isar_Examples/Peirce.thy
author wenzelm
Thu, 30 Oct 2014 22:45:19 +0100
changeset 58839 ccda99401bc8
parent 58614 7338eb25226c
child 58882 6e2010ab8bd9
permissions -rw-r--r--
eliminated aliases;
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
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
     5
header \<open>Peirce's Law\<close>
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
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    11
text \<open>We consider Peirce's Law: $((A \impl B) \impl A) \impl A$.
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    12
  This is an inherently non-intuitionistic statement, so its proof
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    13
  will certainly involve some form of classical contradiction.
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    14
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    15
  The first proof is again a well-balanced combination of plain
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    16
  backward and forward reasoning.  The actual classical step is where
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    17
  the negated goal may be introduced as additional assumption.  This
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    18
  eventually leads to a contradiction.\footnote{The rule involved
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    19
  there is negation elimination; it holds in intuitionistic logic as
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    20
  well.}\<close>
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    21
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    22
theorem "((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    23
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    24
  assume "(A \<longrightarrow> B) \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    25
  show A
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    26
  proof (rule classical)
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    27
    assume "\<not> A"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    28
    have "A \<longrightarrow> B"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    29
    proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    30
      assume A
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    31
      with \<open>\<not> A\<close> show B by contradiction
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    32
    qed
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    33
    with \<open>(A \<longrightarrow> B) \<longrightarrow> A\<close> show A ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    34
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    35
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    36
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    37
text \<open>In the subsequent version the reasoning is rearranged by means
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    38
  of ``weak assumptions'' (as introduced by \isacommand{presume}).
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    39
  Before assuming the negated goal $\neg A$, its intended consequence
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    40
  $A \impl B$ is put into place in order to solve the main problem.
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    41
  Nevertheless, we do not get anything for free, but have to establish
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    42
  $A \impl B$ later on.  The overall effect is that of a logical
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    43
  \emph{cut}.
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    44
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    45
  Technically speaking, whenever some goal is solved by
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    46
  \isacommand{show} in the context of weak assumptions then the latter
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    47
  give rise to new subgoals, which may be established separately.  In
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    48
  contrast, strong assumptions (as introduced by \isacommand{assume})
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    49
  are solved immediately.\<close>
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    50
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    51
theorem "((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    52
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    53
  assume "(A \<longrightarrow> B) \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    54
  show A
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    55
  proof (rule classical)
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    56
    presume "A \<longrightarrow> B"
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    57
    with \<open>(A \<longrightarrow> B) \<longrightarrow> A\<close> show A ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    58
  next
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    59
    assume "\<not> A"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    60
    show "A \<longrightarrow> B"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    61
    proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    62
      assume A
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    63
      with \<open>\<not> A\<close> show B by contradiction
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    64
    qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    65
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    66
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    67
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    68
text \<open>Note that the goals stemming from weak assumptions may be even
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    69
  left until qed time, where they get eventually solved ``by
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    70
  assumption'' as well.  In that case there is really no fundamental
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    71
  difference between the two kinds of assumptions, apart from the
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    72
  order of reducing the individual parts of the proof configuration.
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    73
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    74
  Nevertheless, the ``strong'' mode of plain assumptions is quite
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    75
  important in practice to achieve robustness of proof text
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    76
  interpretation.  By forcing both the conclusion \emph{and} the
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    77
  assumptions to unify with the pending goal to be solved, goal
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    78
  selection becomes quite deterministic.  For example, decomposition
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    79
  with rules of the ``case-analysis'' type usually gives rise to
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    80
  several goals that only differ in there local contexts.  With strong
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    81
  assumptions these may be still solved in any order in a predictable
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
    82
  way, while weak ones would quickly lead to great confusion,
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    83
  eventually demanding even some backtracking.\<close>
7860
7819547df4d8 improved presentation;
wenzelm
parents: 7748
diff changeset
    84
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 7982
diff changeset
    85
end