| author | wenzelm | 
| Tue, 18 Apr 2017 19:17:46 +0200 | |
| changeset 65507 | decdb95bd007 | 
| parent 62279 | f054c364b53f | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 2 | |
| 58889 | 3 | section \<open>Example: First-Order Logic\<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 4 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 5 | theory %visible First_Order_Logic | 
| 42651 | 6 | imports Base (* FIXME Pure!? *) | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 7 | begin | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 8 | |
| 58618 | 9 | text \<open> | 
| 62279 | 10 | In order to commence a new object-logic within Isabelle/Pure we introduce | 
| 11 | abstract syntactic categories \<open>i\<close> for individuals and \<open>o\<close> for | |
| 12 | object-propositions. The latter is embedded into the language of Pure | |
| 13 | propositions by means of a separate judgment. | |
| 58618 | 14 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 15 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 16 | typedecl i | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 17 | typedecl o | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 18 | |
| 62279 | 19 | judgment Trueprop :: "o \<Rightarrow> prop"    ("_" 5)
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 20 | |
| 58618 | 21 | text \<open> | 
| 62279 | 22 | Note that the object-logic judgment is implicit in the syntax: writing | 
| 23 |   @{prop A} produces @{term "Trueprop A"} internally. From the Pure
 | |
| 24 |   perspective this means ``@{prop A} is derivable in the object-logic''.
 | |
| 58618 | 25 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 26 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 27 | |
| 58618 | 28 | subsection \<open>Equational reasoning \label{sec:framework-ex-equal}\<close>
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 29 | |
| 58618 | 30 | text \<open> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 31 | Equality is axiomatized as a binary predicate on individuals, with | 
| 62279 | 32 | reflexivity as introduction, and substitution as elimination principle. Note | 
| 33 | that the latter is particularly convenient in a framework like Isabelle, | |
| 34 | because syntactic congruences are implicitly produced by unification of | |
| 35 | \<open>B x\<close> against expressions containing occurrences of \<open>x\<close>. | |
| 58618 | 36 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 37 | |
| 62279 | 38 | axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infix "=" 50) | 
| 39 | where refl [intro]: "x = x" | |
| 40 | and subst [elim]: "x = y \<Longrightarrow> B x \<Longrightarrow> B y" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 41 | |
| 58618 | 42 | text \<open> | 
| 62279 | 43 | Substitution is very powerful, but also hard to control in full generality. | 
| 44 |   We derive some common symmetry~/ transitivity schemes of @{term equal} as
 | |
| 45 | particular consequences. | |
| 58618 | 46 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 47 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 48 | theorem sym [sym]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 49 | assumes "x = y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 50 | shows "y = x" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 51 | proof - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 52 | have "x = x" .. | 
| 58618 | 53 | with \<open>x = y\<close> show "y = x" .. | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 54 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 55 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 56 | theorem forw_subst [trans]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 57 | assumes "y = x" and "B x" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 58 | shows "B y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 59 | proof - | 
| 58618 | 60 | from \<open>y = x\<close> have "x = y" .. | 
| 61 | from this and \<open>B x\<close> show "B y" .. | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 62 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 63 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 64 | theorem back_subst [trans]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 65 | assumes "B x" and "x = y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 66 | shows "B y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 67 | proof - | 
| 58618 | 68 | from \<open>x = y\<close> and \<open>B x\<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 69 | show "B y" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 70 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 71 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 72 | theorem trans [trans]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 73 | assumes "x = y" and "y = z" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 74 | shows "x = z" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 75 | proof - | 
| 58618 | 76 | from \<open>y = z\<close> and \<open>x = y\<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 77 | show "x = z" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 78 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 79 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 80 | |
| 58618 | 81 | subsection \<open>Basic group theory\<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 82 | |
| 58618 | 83 | text \<open> | 
| 62279 | 84 | As an example for equational reasoning we consider some bits of group | 
| 85 | theory. The subsequent locale definition postulates group operations and | |
| 86 | axioms; we also derive some consequences of this specification. | |
| 58618 | 87 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 88 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 89 | locale group = | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 90 | fixes prod :: "i \<Rightarrow> i \<Rightarrow> i" (infix "\<circ>" 70) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 91 |     and inv :: "i \<Rightarrow> i"  ("(_\<inverse>)" [1000] 999)
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 92 |     and unit :: i  ("1")
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 93 | assumes assoc: "(x \<circ> y) \<circ> z = x \<circ> (y \<circ> z)" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 94 | and left_unit: "1 \<circ> x = x" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 95 | and left_inv: "x\<inverse> \<circ> x = 1" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 96 | begin | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 97 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 98 | theorem right_inv: "x \<circ> x\<inverse> = 1" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 99 | proof - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 100 | have "x \<circ> x\<inverse> = 1 \<circ> (x \<circ> x\<inverse>)" by (rule left_unit [symmetric]) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 101 | also have "\<dots> = (1 \<circ> x) \<circ> x\<inverse>" by (rule assoc [symmetric]) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 102 | also have "1 = (x\<inverse>)\<inverse> \<circ> x\<inverse>" by (rule left_inv [symmetric]) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 103 | also have "\<dots> \<circ> x = (x\<inverse>)\<inverse> \<circ> (x\<inverse> \<circ> x)" by (rule assoc) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 104 | also have "x\<inverse> \<circ> x = 1" by (rule left_inv) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 105 | also have "((x\<inverse>)\<inverse> \<circ> \<dots>) \<circ> x\<inverse> = (x\<inverse>)\<inverse> \<circ> (1 \<circ> x\<inverse>)" by (rule assoc) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 106 | also have "1 \<circ> x\<inverse> = x\<inverse>" by (rule left_unit) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 107 | also have "(x\<inverse>)\<inverse> \<circ> \<dots> = 1" by (rule left_inv) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 108 | finally show "x \<circ> x\<inverse> = 1" . | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 109 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 110 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 111 | theorem right_unit: "x \<circ> 1 = x" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 112 | proof - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 113 | have "1 = x\<inverse> \<circ> x" by (rule left_inv [symmetric]) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 114 | also have "x \<circ> \<dots> = (x \<circ> x\<inverse>) \<circ> x" by (rule assoc [symmetric]) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 115 | also have "x \<circ> x\<inverse> = 1" by (rule right_inv) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 116 | also have "\<dots> \<circ> x = x" by (rule left_unit) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 117 | finally show "x \<circ> 1 = x" . | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 118 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 119 | |
| 58618 | 120 | text \<open> | 
| 62279 | 121 | Reasoning from basic axioms is often tedious. Our proofs work by producing | 
| 122 | various instances of the given rules (potentially the symmetric form) using | |
| 123 | the pattern ``\<^theory_text>\<open>have eq by (rule r)\<close>'' and composing the chain of results | |
| 124 | via \<^theory_text>\<open>also\<close>/\<^theory_text>\<open>finally\<close>. These steps may involve any of the transitivity | |
| 125 |   rules declared in \secref{sec:framework-ex-equal}, namely @{thm trans} in
 | |
| 126 |   combining the first two results in @{thm right_inv} and in the final steps
 | |
| 127 |   of both proofs, @{thm forw_subst} in the first combination of @{thm
 | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 128 |   right_unit}, and @{thm back_subst} in all other calculational steps.
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 129 | |
| 62279 | 130 | Occasional substitutions in calculations are adequate, but should not be | 
| 131 | over-emphasized. The other extreme is to compose a chain by plain | |
| 132 | transitivity only, with replacements occurring always in topmost position. | |
| 133 | For example: | |
| 58618 | 134 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 135 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 136 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 137 | theorem "\<And>A. PROP A \<Longrightarrow> PROP A" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 138 | proof - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 139 | assume [symmetric, defn]: "\<And>x y. (x \<equiv> y) \<equiv> Trueprop (x = y)" | 
| 56594 | 140 | fix x | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 141 | (*>*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 142 | have "x \<circ> 1 = x \<circ> (x\<inverse> \<circ> x)" unfolding left_inv .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 143 | also have "\<dots> = (x \<circ> x\<inverse>) \<circ> x" unfolding assoc .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 144 | also have "\<dots> = 1 \<circ> x" unfolding right_inv .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 145 | also have "\<dots> = x" unfolding left_unit .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 146 | finally have "x \<circ> 1 = x" . | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 147 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 148 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 149 | (*>*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 150 | |
| 58618 | 151 | text \<open> | 
| 62279 | 152 | Here we have re-used the built-in mechanism for unfolding definitions in | 
| 153 | order to normalize each equational problem. A more realistic object-logic | |
| 154 |   would include proper setup for the Simplifier (\secref{sec:simplifier}), the
 | |
| 155 | main automated tool for equational reasoning in Isabelle. Then ``\<^theory_text>\<open>unfolding | |
| 156 | left_inv ..\<close>'' would become ``\<^theory_text>\<open>by (simp only: left_inv)\<close>'' etc. | |
| 58618 | 157 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 158 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 159 | end | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 160 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 161 | |
| 58618 | 162 | subsection \<open>Propositional logic \label{sec:framework-ex-prop}\<close>
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 163 | |
| 58618 | 164 | text \<open> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 165 | We axiomatize basic connectives of propositional logic: implication, | 
| 62279 | 166 | disjunction, and conjunction. The associated rules are modeled after | 
| 167 |   Gentzen's system of Natural Deduction @{cite "Gentzen:1935"}.
 | |
| 58618 | 168 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 169 | |
| 62279 | 170 | axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25) | 
| 171 | where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B" | |
| 172 | and impD [dest]: "(A \<longrightarrow> B) \<Longrightarrow> A \<Longrightarrow> B" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 173 | |
| 62279 | 174 | axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) | 
| 175 | where disjI\<^sub>1 [intro]: "A \<Longrightarrow> A \<or> B" | |
| 176 | and disjI\<^sub>2 [intro]: "B \<Longrightarrow> A \<or> B" | |
| 177 | and disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 178 | |
| 62279 | 179 | axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) | 
| 180 | where conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B" | |
| 181 | and conjD\<^sub>1: "A \<and> B \<Longrightarrow> A" | |
| 182 | and conjD\<^sub>2: "A \<and> B \<Longrightarrow> B" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 183 | |
| 58618 | 184 | text \<open> | 
| 62279 | 185 |   The conjunctive destructions have the disadvantage that decomposing @{prop
 | 
| 186 | "A \<and> B"} involves an immediate decision which component should be projected. | |
| 187 |   The more convenient simultaneous elimination @{prop "A \<and> B \<Longrightarrow> (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow>
 | |
| 188 | C"} can be derived as follows: | |
| 58618 | 189 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 190 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 191 | theorem conjE [elim]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 192 | assumes "A \<and> B" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 193 | obtains A and B | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 194 | proof | 
| 58618 | 195 | from \<open>A \<and> B\<close> show A by (rule conjD\<^sub>1) | 
| 196 | from \<open>A \<and> B\<close> show B by (rule conjD\<^sub>2) | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 197 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 198 | |
| 58618 | 199 | text \<open> | 
| 62279 | 200 | Here is an example of swapping conjuncts with a single intermediate | 
| 201 | elimination step: | |
| 58618 | 202 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 203 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 204 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 205 | lemma "\<And>A. PROP A \<Longrightarrow> PROP A" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 206 | proof - | 
| 62279 | 207 | fix A B | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 208 | (*>*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 209 | assume "A \<and> B" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 210 | then obtain B and A .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 211 | then have "B \<and> A" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 212 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 213 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 214 | (*>*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 215 | |
| 58618 | 216 | text \<open> | 
| 62279 | 217 | Note that the analogous elimination rule for disjunction ``\<^theory_text>\<open>assumes "A \<or> B" | 
| 218 |   obtains A \<BBAR> B\<close>'' coincides with the original axiomatization of @{thm
 | |
| 219 | disjE}. | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 220 | |
| 61421 | 221 | \<^medskip> | 
| 62279 | 222 | We continue propositional logic by introducing absurdity with its | 
| 223 | characteristic elimination. Plain truth may then be defined as a proposition | |
| 224 | that is trivially true. | |
| 58618 | 225 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 226 | |
| 62279 | 227 | axiomatization false :: o  ("\<bottom>")
 | 
| 228 | where falseE [elim]: "\<bottom> \<Longrightarrow> A" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 229 | |
| 62279 | 230 | definition true :: o  ("\<top>")
 | 
| 231 | where "\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 232 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 233 | theorem trueI [intro]: \<top> | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 234 | unfolding true_def .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 235 | |
| 58618 | 236 | text \<open> | 
| 61421 | 237 | \<^medskip> | 
| 238 | Now negation represents an implication towards absurdity: | |
| 58618 | 239 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 240 | |
| 62279 | 241 | definition not :: "o \<Rightarrow> o"  ("\<not> _" [40] 40)
 | 
| 242 | where "\<not> A \<equiv> A \<longrightarrow> \<bottom>" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 243 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 244 | theorem notI [intro]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 245 | assumes "A \<Longrightarrow> \<bottom>" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 246 | shows "\<not> A" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 247 | unfolding not_def | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 248 | proof | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 249 | assume A | 
| 58618 | 250 | then show \<bottom> by (rule \<open>A \<Longrightarrow> \<bottom>\<close>) | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 251 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 252 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 253 | theorem notE [elim]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 254 | assumes "\<not> A" and A | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 255 | shows B | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 256 | proof - | 
| 58618 | 257 | from \<open>\<not> A\<close> have "A \<longrightarrow> \<bottom>" unfolding not_def . | 
| 258 | from \<open>A \<longrightarrow> \<bottom>\<close> and \<open>A\<close> have \<bottom> .. | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 259 | then show B .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 260 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 261 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 262 | |
| 58618 | 263 | subsection \<open>Classical logic\<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 264 | |
| 58618 | 265 | text \<open> | 
| 62279 | 266 | Subsequently we state the principle of classical contradiction as a local | 
| 267 | assumption. Thus we refrain from forcing the object-logic into the classical | |
| 268 | perspective. Within that context, we may derive well-known consequences of | |
| 269 | the classical principle. | |
| 58618 | 270 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 271 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 272 | locale classical = | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 273 | assumes classical: "(\<not> C \<Longrightarrow> C) \<Longrightarrow> C" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 274 | begin | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 275 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 276 | theorem double_negation: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 277 | assumes "\<not> \<not> C" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 278 | shows C | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 279 | proof (rule classical) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 280 | assume "\<not> C" | 
| 58618 | 281 | with \<open>\<not> \<not> C\<close> show C .. | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 282 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 283 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 284 | theorem tertium_non_datur: "C \<or> \<not> C" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 285 | proof (rule double_negation) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 286 | show "\<not> \<not> (C \<or> \<not> C)" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 287 | proof | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 288 | assume "\<not> (C \<or> \<not> C)" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 289 | have "\<not> C" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 290 | proof | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 291 | assume C then have "C \<or> \<not> C" .. | 
| 58618 | 292 | with \<open>\<not> (C \<or> \<not> C)\<close> show \<bottom> .. | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 293 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 294 | then have "C \<or> \<not> C" .. | 
| 58618 | 295 | with \<open>\<not> (C \<or> \<not> C)\<close> show \<bottom> .. | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 296 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 297 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 298 | |
| 58618 | 299 | text \<open> | 
| 62279 | 300 | These examples illustrate both classical reasoning and non-trivial | 
| 301 | propositional proofs in general. All three rules characterize classical | |
| 302 | logic independently, but the original rule is already the most convenient to | |
| 303 |   use, because it leaves the conclusion unchanged. Note that @{prop "(\<not> C \<Longrightarrow> C)
 | |
| 304 | \<Longrightarrow> C"} fits again into our format for eliminations, despite the additional | |
| 305 |   twist that the context refers to the main conclusion. So we may write @{thm
 | |
| 306 | classical} as the Isar statement ``\<^theory_text>\<open>obtains \<not> thesis\<close>''. This also explains | |
| 307 | nicely how classical reasoning really works: whatever the main \<open>thesis\<close> | |
| 308 | might be, we may always assume its negation! | |
| 58618 | 309 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 310 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 311 | end | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 312 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 313 | |
| 58618 | 314 | subsection \<open>Quantifiers \label{sec:framework-ex-quant}\<close>
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 315 | |
| 58618 | 316 | text \<open> | 
| 62279 | 317 | Representing quantifiers is easy, thanks to the higher-order nature of the | 
| 318 | underlying framework. According to the well-known technique introduced by | |
| 319 |   Church @{cite "church40"}, quantifiers are operators on predicates, which
 | |
| 320 |   are syntactically represented as \<open>\<lambda>\<close>-terms of type @{typ "i \<Rightarrow> o"}. Binder
 | |
| 321 | notation turns \<open>All (\<lambda>x. B x)\<close> into \<open>\<forall>x. B x\<close> etc. | |
| 58618 | 322 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 323 | |
| 62279 | 324 | axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10) | 
| 325 | where allI [intro]: "(\<And>x. B x) \<Longrightarrow> \<forall>x. B x" | |
| 326 | and allD [dest]: "(\<forall>x. B x) \<Longrightarrow> B a" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 327 | |
| 62279 | 328 | axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10) | 
| 329 | where exI [intro]: "B a \<Longrightarrow> (\<exists>x. B x)" | |
| 330 | and exE [elim]: "(\<exists>x. B x) \<Longrightarrow> (\<And>x. B x \<Longrightarrow> C) \<Longrightarrow> C" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 331 | |
| 58618 | 332 | text \<open> | 
| 62279 | 333 |   The statement of @{thm exE} corresponds to ``\<^theory_text>\<open>assumes "\<exists>x. B x" obtains x
 | 
| 334 | where "B x"\<close>'' in Isar. In the subsequent example we illustrate quantifier | |
| 335 | reasoning involving all four rules: | |
| 58618 | 336 | \<close> | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 337 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 338 | theorem | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 339 | assumes "\<exists>x. \<forall>y. R x y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 340 | shows "\<forall>y. \<exists>x. R x y" | 
| 61580 | 341 | proof \<comment> \<open>\<open>\<forall>\<close> introduction\<close> | 
| 342 | obtain x where "\<forall>y. R x y" using \<open>\<exists>x. \<forall>y. R x y\<close> .. \<comment> \<open>\<open>\<exists>\<close> elimination\<close> | |
| 343 | fix y have "R x y" using \<open>\<forall>y. R x y\<close> .. \<comment> \<open>\<open>\<forall>\<close> destruction\<close> | |
| 344 | then show "\<exists>x. R x y" .. \<comment> \<open>\<open>\<exists>\<close> introduction\<close> | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 345 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 346 | |
| 29734 | 347 | |
| 58618 | 348 | subsection \<open>Canonical reasoning patterns\<close> | 
| 29734 | 349 | |
| 58618 | 350 | text \<open> | 
| 29734 | 351 | The main rules of first-order predicate logic from | 
| 62279 | 352 |   \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} can now
 | 
| 353 | be summarized as follows, using the native Isar statement format of | |
| 354 |   \secref{sec:framework-stmt}.
 | |
| 29734 | 355 | |
| 61421 | 356 | \<^medskip> | 
| 29734 | 357 |   \begin{tabular}{l}
 | 
| 62279 | 358 | \<^theory_text>\<open>impI: assumes "A \<Longrightarrow> B" shows "A \<longrightarrow> B"\<close> \\ | 
| 359 | \<^theory_text>\<open>impD: assumes "A \<longrightarrow> B" and A shows B\<close> \\[1ex] | |
| 29734 | 360 | |
| 62279 | 361 | \<^theory_text>\<open>disjI\<^sub>1: assumes A shows "A \<or> B"\<close> \\ | 
| 362 | \<^theory_text>\<open>disjI\<^sub>2: assumes B shows "A \<or> B"\<close> \\ | |
| 363 | \<^theory_text>\<open>disjE: assumes "A \<or> B" obtains A \<BBAR> B\<close> \\[1ex] | |
| 29734 | 364 | |
| 62279 | 365 | \<^theory_text>\<open>conjI: assumes A and B shows A \<and> B\<close> \\ | 
| 366 | \<^theory_text>\<open>conjE: assumes "A \<and> B" obtains A and B\<close> \\[1ex] | |
| 29734 | 367 | |
| 62279 | 368 | \<^theory_text>\<open>falseE: assumes \<bottom> shows A\<close> \\ | 
| 369 | \<^theory_text>\<open>trueI: shows \<top>\<close> \\[1ex] | |
| 29734 | 370 | |
| 62279 | 371 | \<^theory_text>\<open>notI: assumes "A \<Longrightarrow> \<bottom>" shows "\<not> A"\<close> \\ | 
| 372 | \<^theory_text>\<open>notE: assumes "\<not> A" and A shows B\<close> \\[1ex] | |
| 29734 | 373 | |
| 62279 | 374 | \<^theory_text>\<open>allI: assumes "\<And>x. B x" shows "\<forall>x. B x"\<close> \\ | 
| 375 | \<^theory_text>\<open>allE: assumes "\<forall>x. B x" shows "B a"\<close> \\[1ex] | |
| 29734 | 376 | |
| 62279 | 377 | \<^theory_text>\<open>exI: assumes "B a" shows "\<exists>x. B x"\<close> \\ | 
| 378 | \<^theory_text>\<open>exE: assumes "\<exists>x. B x" obtains a where "B a"\<close> | |
| 29734 | 379 |   \end{tabular}
 | 
| 61421 | 380 | \<^medskip> | 
| 29734 | 381 | |
| 62279 | 382 | This essentially provides a declarative reading of Pure rules as Isar | 
| 383 | reasoning patterns: the rule statements tells how a canonical proof outline | |
| 384 | shall look like. Since the above rules have already been declared as | |
| 385 |   @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure)
 | |
| 386 | dest} --- each according to its particular shape --- we can immediately | |
| 387 | write Isar proof texts as follows: | |
| 58618 | 388 | \<close> | 
| 29734 | 389 | |
| 390 | (*<*) | |
| 391 | theorem "\<And>A. PROP A \<Longrightarrow> PROP A" | |
| 392 | proof - | |
| 393 | (*>*) | |
| 394 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 395 |   text_raw \<open>\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 396 | |
| 397 | have "A \<longrightarrow> B" | |
| 398 | proof | |
| 399 | assume A | |
| 62271 | 400 | show B \<proof> | 
| 29734 | 401 | qed | 
| 402 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 403 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 404 | |
| 62271 | 405 | have "A \<longrightarrow> B" and A \<proof> | 
| 29734 | 406 | then have B .. | 
| 407 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 408 |   text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 409 | |
| 62271 | 410 | have A \<proof> | 
| 29734 | 411 | then have "A \<or> B" .. | 
| 412 | ||
| 62271 | 413 | have B \<proof> | 
| 29734 | 414 | then have "A \<or> B" .. | 
| 415 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 416 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 417 | |
| 62271 | 418 | have "A \<or> B" \<proof> | 
| 29734 | 419 | then have C | 
| 420 | proof | |
| 421 | assume A | |
| 62271 | 422 | then show C \<proof> | 
| 29734 | 423 | next | 
| 424 | assume B | |
| 62271 | 425 | then show C \<proof> | 
| 29734 | 426 | qed | 
| 427 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 428 |   text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 429 | |
| 62271 | 430 | have A and B \<proof> | 
| 29734 | 431 | then have "A \<and> B" .. | 
| 432 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 433 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 434 | |
| 62271 | 435 | have "A \<and> B" \<proof> | 
| 29734 | 436 | then obtain A and B .. | 
| 437 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 438 |   text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 439 | |
| 62271 | 440 | have "\<bottom>" \<proof> | 
| 29734 | 441 | then have A .. | 
| 442 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 443 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 444 | |
| 445 | have "\<top>" .. | |
| 446 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 447 |   text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 448 | |
| 449 | have "\<not> A" | |
| 450 | proof | |
| 451 | assume A | |
| 62271 | 452 | then show "\<bottom>" \<proof> | 
| 29734 | 453 | qed | 
| 454 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 455 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 456 | |
| 62271 | 457 | have "\<not> A" and A \<proof> | 
| 29734 | 458 | then have B .. | 
| 459 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 460 |   text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 461 | |
| 462 | have "\<forall>x. B x" | |
| 463 | proof | |
| 464 | fix x | |
| 62271 | 465 | show "B x" \<proof> | 
| 29734 | 466 | qed | 
| 467 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 468 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 469 | |
| 62271 | 470 | have "\<forall>x. B x" \<proof> | 
| 29734 | 471 | then have "B a" .. | 
| 472 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 473 |   text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 474 | |
| 475 | have "\<exists>x. B x" | |
| 476 | proof | |
| 62271 | 477 | show "B a" \<proof> | 
| 29734 | 478 | qed | 
| 479 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 480 |   text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 481 | |
| 62271 | 482 | have "\<exists>x. B x" \<proof> | 
| 29734 | 483 | then obtain a where "B a" .. | 
| 484 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58889diff
changeset | 485 |   text_raw \<open>\end{minipage}\<close>
 | 
| 29734 | 486 | |
| 487 | (*<*) | |
| 488 | qed | |
| 489 | (*>*) | |
| 490 | ||
| 58618 | 491 | text \<open> | 
| 61421 | 492 | \<^bigskip> | 
| 62279 | 493 | Of course, these proofs are merely examples. As sketched in | 
| 494 |   \secref{sec:framework-subproof}, there is a fair amount of flexibility in
 | |
| 495 | expressing Pure deductions in Isar. Here the user is asked to express | |
| 496 | himself adequately, aiming at proof texts of literary quality. | |
| 58618 | 497 | \<close> | 
| 29734 | 498 | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 499 | end %visible |