| author | wenzelm | 
| Sun, 29 Jul 2012 21:40:46 +0200 | |
| changeset 48592 | a125b8040ada | 
| parent 39157 | b98909faaea8 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
32960diff
changeset | 1 | (* Title: HOL/Proofs/ex/Hilbert_Classical.thy | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 2 | Author: Stefan Berghofer and Markus Wenzel, TU Muenchen | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 3 | *) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 4 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 5 | header {* Hilbert's choice and classical logic *}
 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 6 | |
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
32960diff
changeset | 7 | theory Hilbert_Classical | 
| 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
32960diff
changeset | 8 | imports Main | 
| 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
32960diff
changeset | 9 | begin | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 10 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 11 | text {*
 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 12 | Derivation of the classical law of tertium-non-datur by means of | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 13 | Hilbert's choice operator (due to M. J. Beeson and J. Harrison). | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 14 | *} | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 15 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 16 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 17 | subsection {* Proof text *}
 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 18 | |
| 11591 | 19 | theorem tnd: "A \<or> \<not> A" | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 20 | proof - | 
| 11590 | 21 | let ?P = "\<lambda>X. X = False \<or> X = True \<and> A" | 
| 22 | let ?Q = "\<lambda>X. X = False \<and> A \<or> X = True" | |
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 23 | |
| 11590 | 24 | have a: "?P (Eps ?P)" | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 25 | proof (rule someI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 26 | have "False = False" .. | 
| 11590 | 27 | thus "?P False" .. | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 28 | qed | 
| 11590 | 29 | have b: "?Q (Eps ?Q)" | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 30 | proof (rule someI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 31 | have "True = True" .. | 
| 11590 | 32 | thus "?Q True" .. | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 33 | qed | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 34 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 35 | from a show ?thesis | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 36 | proof | 
| 11590 | 37 | assume "Eps ?P = True \<and> A" | 
| 38 | hence A .. | |
| 39 | thus ?thesis .. | |
| 40 | next | |
| 41 | assume P: "Eps ?P = False" | |
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 42 | from b show ?thesis | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 43 | proof | 
| 11590 | 44 | assume "Eps ?Q = False \<and> A" | 
| 45 | hence A .. | |
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 46 | thus ?thesis .. | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 47 | next | 
| 11590 | 48 | assume Q: "Eps ?Q = True" | 
| 49 | have neq: "?P \<noteq> ?Q" | |
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 50 | proof | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 51 | assume "?P = ?Q" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 52 | hence "Eps ?P = Eps ?Q" by (rule arg_cong) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 53 | also note P | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 54 | also note Q | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 55 | finally show False by (rule False_neq_True) | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 56 | qed | 
| 11590 | 57 | have "\<not> A" | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 58 | proof | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 59 | assume a: A | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 60 | have "?P = ?Q" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 61 | proof (rule ext) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 62 | fix x show "?P x = ?Q x" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 63 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 64 | assume "?P x" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 65 | thus "?Q x" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 66 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 67 | assume "x = False" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 68 | from this and a have "x = False \<and> A" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 69 | thus "?Q x" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 70 | next | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 71 | assume "x = True \<and> A" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 72 | hence "x = True" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 73 | thus "?Q x" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 74 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 75 | next | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 76 | assume "?Q x" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 77 | thus "?P x" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 78 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 79 | assume "x = False \<and> A" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 80 | hence "x = False" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 81 | thus "?P x" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 82 | next | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 83 | assume "x = True" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 84 | from this and a have "x = True \<and> A" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 85 | thus "?P x" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 86 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 87 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 88 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
26828diff
changeset | 89 | with neq show False by contradiction | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 90 | qed | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 91 | thus ?thesis .. | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 92 | qed | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 93 | qed | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 94 | qed | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 95 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 96 | |
| 11591 | 97 | subsection {* Proof term of text *}
 | 
| 98 | ||
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
32960diff
changeset | 99 | prf tnd | 
| 11591 | 100 | |
| 101 | ||
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 102 | subsection {* Proof script *}
 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 103 | |
| 11590 | 104 | theorem tnd': "A \<or> \<not> A" | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 105 | apply (subgoal_tac | 
| 11590 | 106 | "(((SOME x. x = False \<or> x = True \<and> A) = False) \<or> | 
| 107 | ((SOME x. x = False \<or> x = True \<and> A) = True) \<and> A) \<and> | |
| 108 | (((SOME x. x = False \<and> A \<or> x = True) = False) \<and> A \<or> | |
| 109 | ((SOME x. x = False \<and> A \<or> x = True) = True))") | |
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 110 | prefer 2 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 111 | apply (rule conjI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 112 | apply (rule someI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 113 | apply (rule disjI1) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 114 | apply (rule refl) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 115 | apply (rule someI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 116 | apply (rule disjI2) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 117 | apply (rule refl) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 118 | apply (erule conjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 119 | apply (erule disjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 120 | apply (erule disjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 121 | apply (erule conjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 122 | apply (erule disjI1) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 123 | prefer 2 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 124 | apply (erule conjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 125 | apply (erule disjI1) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 126 | apply (subgoal_tac | 
| 11590 | 127 | "(\<lambda>x. (x = False) \<or> (x = True) \<and> A) \<noteq> | 
| 128 | (\<lambda>x. (x = False) \<and> A \<or> (x = True))") | |
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 129 | prefer 2 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 130 | apply (rule notI) | 
| 11590 | 131 | apply (drule_tac f = "\<lambda>y. SOME x. y x" in arg_cong) | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 132 | apply (drule trans, assumption) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 133 | apply (drule sym) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 134 | apply (drule trans, assumption) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 135 | apply (erule False_neq_True) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 136 | apply (rule disjI2) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 137 | apply (rule notI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 138 | apply (erule notE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 139 | apply (rule ext) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 140 | apply (rule iffI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 141 | apply (erule disjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 142 | apply (rule disjI1) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 143 | apply (erule conjI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 144 | apply assumption | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 145 | apply (erule conjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 146 | apply (erule disjI2) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 147 | apply (erule disjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 148 | apply (erule conjE) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 149 | apply (erule disjI1) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 150 | apply (rule disjI2) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 151 | apply (erule conjI) | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 152 | apply assumption | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 153 | done | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 154 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 155 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 156 | subsection {* Proof term of script *}
 | 
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 157 | |
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
32960diff
changeset | 158 | prf tnd' | 
| 11584 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 159 | |
| 
c8e98b9498b4
derive tertium-non-datur by means of Hilbert's choice operator;
 wenzelm parents: diff
changeset | 160 | end |