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