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