| author | krauss | 
| Sun, 21 Aug 2011 22:13:04 +0200 | |
| changeset 44367 | 74c08021ab2e | 
| parent 42651 | e3fdb7c96be5 | 
| child 45103 | a45121ffcfcb | 
| permissions | -rw-r--r-- | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 1 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 2 | header {* Example: First-Order Logic *}
 | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 8 | text {*
 | 
| 29732 | 9 | \noindent In order to commence a new object-logic within | 
| 10 |   Isabelle/Pure we introduce abstract syntactic categories @{text "i"}
 | |
| 11 |   for individuals and @{text "o"} for object-propositions.  The latter
 | |
| 12 | is embedded into the language of Pure propositions by means of a | |
| 13 | separate judgment. | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 14 | *} | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 22 | text {*
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 23 | \noindent Note that the object-logic judgement is implicit in the | 
| 
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''. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 27 | *} | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 30 | subsection {* Equational reasoning \label{sec:framework-ex-equal} *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 31 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 32 | text {*
 | 
| 
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}.
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 39 | *} | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 47 | text {*
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 48 | \noindent Substitution is very powerful, but also hard to control in | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 49 | full generality. We derive some common symmetry~/ transitivity | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 50 | schemes of as particular consequences. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 51 | *} | 
| 
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" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 58 | with `x = y` show "y = x" .. | 
| 
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 - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 65 | from `y = x` have "x = y" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 66 | from this and `B x` show "B y" .. | 
| 
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 - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 73 | from `x = y` and `B x` | 
| 
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 - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 81 | from `y = z` and `x = y` | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 86 | subsection {* Basic group theory *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 87 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 88 | text {*
 | 
| 
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. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 93 | *} | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 126 | text {*
 | 
| 29732 | 127 | \noindent Reasoning from basic axioms is often tedious. Our proofs | 
| 128 | work by producing various instances of the given rules (potentially | |
| 129 |   the symmetric form) using the pattern ``@{command have}~@{text
 | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 130 |   eq}~@{command "by"}~@{text "(rule r)"}'' and composing the chain of
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 131 |   results via @{command also}/@{command finally}.  These steps may
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 132 | involve any of the transitivity rules declared in | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 133 |   \secref{sec:framework-ex-equal}, namely @{thm trans} in combining
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 134 |   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 | 135 |   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 | 136 |   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 | 137 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 138 | Occasional substitutions in calculations are adequate, but should | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 139 | 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 | 140 | plain transitivity only, with replacements occurring always in | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 141 | topmost position. For example: | 
| 
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 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 145 | theorem "\<And>A. PROP A \<Longrightarrow> PROP A" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 146 | proof - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 147 | assume [symmetric, defn]: "\<And>x y. (x \<equiv> y) \<equiv> Trueprop (x = y)" | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 158 | text {*
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 159 | \noindent Here we have re-used the built-in mechanism for unfolding | 
| 
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
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 164 |   left_inv}~@{command ".."}'' would become ``@{command "by"}~@{text
 | 
| 29735 | 165 | "(simp only: left_inv)"}'' etc. | 
| 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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 168 | end | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 169 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 170 | |
| 29734 | 171 | subsection {* Propositional logic \label{sec:framework-ex-prop} *}
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 172 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 173 | text {*
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 174 | We axiomatize basic connectives of propositional logic: implication, | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 175 | disjunction, and conjunction. The associated rules are modeled | 
| 29732 | 176 |   after Gentzen's system of Natural Deduction \cite{Gentzen:1935}.
 | 
| 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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 179 | axiomatization | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 180 | imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25) where | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 181 | impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B" and | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 182 | impD [dest]: "(A \<longrightarrow> B) \<Longrightarrow> A \<Longrightarrow> B" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 183 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 184 | axiomatization | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 185 | disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) where | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 186 | disjI\<^isub>1 [intro]: "A \<Longrightarrow> A \<or> B" and | 
| 29734 | 187 | disjI\<^isub>2 [intro]: "B \<Longrightarrow> A \<or> B" and | 
| 188 | 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 | 189 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 190 | axiomatization | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 191 | conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) where | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 192 | conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B" and | 
| 29734 | 193 | conjD\<^isub>1: "A \<and> B \<Longrightarrow> A" and | 
| 194 | conjD\<^isub>2: "A \<and> B \<Longrightarrow> B" | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 195 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 196 | text {*
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 197 | \noindent The conjunctive destructions have the disadvantage that | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 198 |   decomposing @{prop "A \<and> B"} involves an immediate decision which
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 199 | component should be projected. The more convenient simultaneous | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 200 |   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 | 201 | follows: | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 204 | theorem conjE [elim]: | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 205 | assumes "A \<and> B" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 206 | obtains A and B | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 207 | proof | 
| 29734 | 208 | from `A \<and> B` show A by (rule conjD\<^isub>1) | 
| 209 | from `A \<and> B` show B by (rule conjD\<^isub>2) | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 210 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 211 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 212 | text {*
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 213 | \noindent Here is an example of swapping conjuncts with a single | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 214 | intermediate elimination step: | 
| 
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 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 218 | lemma "\<And>A. PROP A \<Longrightarrow> PROP A" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 219 | proof - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 220 | (*>*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 221 | assume "A \<and> B" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 222 | then obtain B and A .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 223 | then have "B \<and> A" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 224 | (*<*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 225 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 226 | (*>*) | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 227 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 228 | text {*
 | 
| 29732 | 229 | \noindent Note that the analogous elimination rule for disjunction | 
| 230 |   ``@{text "\<ASSUMES> A \<or> B \<OBTAINS> A \<BBAR> B"}'' coincides with
 | |
| 231 |   the original axiomatization of @{thm disjE}.
 | |
| 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 | \medskip We continue propositional logic by introducing absurdity | 
| 
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. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 236 | *} | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 249 | text {*
 | 
| 29732 | 250 | \medskip\noindent Now negation represents an implication towards | 
| 251 | absurdity: | |
| 29730 
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 | |
| 
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 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 264 | then show \<bottom> by (rule `A \<Longrightarrow> \<bottom>`) | 
| 
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 - | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 271 | from `\<not> A` have "A \<longrightarrow> \<bottom>" unfolding not_def . | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 272 | from `A \<longrightarrow> \<bottom>` and `A` have \<bottom> .. | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 277 | subsection {* Classical logic *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 278 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 279 | text {*
 | 
| 
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. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 284 | *} | 
| 
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" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 295 | with `\<not> \<not> C` show C .. | 
| 
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" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 306 | with `\<not> (C \<or> \<not> C)` show \<bottom> .. | 
| 
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" .. | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 309 | with `\<not> (C \<or> \<not> C)` show \<bottom> .. | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 313 | text {*
 | 
| 29735 | 314 | \noindent These examples illustrate both classical reasoning and | 
| 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
 | |
| 321 |   classical} as the Isar statement ``@{text "\<OBTAINS> \<not> thesis"}''.
 | |
| 322 | This also explains nicely how classical reasoning really works: | |
| 323 |   whatever the main @{text thesis} might be, we may always assume its
 | |
| 324 | negation! | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 325 | *} | 
| 
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 | |
| 29734 | 330 | subsection {* Quantifiers \label{sec:framework-ex-quant} *}
 | 
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 331 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 332 | text {*
 | 
| 
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 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 335 |   introduced by Church \cite{church40}, quantifiers are operators on
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 336 |   predicates, which are syntactically represented as @{text "\<lambda>"}-terms
 | 
| 29732 | 337 |   of type @{typ "i \<Rightarrow> o"}.  Binder notation turns @{text "All (\<lambda>x. B
 | 
| 338 |   x)"} into @{text "\<forall>x. B x"} etc.
 | |
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 339 | *} | 
| 
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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 351 | text {*
 | 
| 29735 | 352 |   \noindent The statement of @{thm exE} corresponds to ``@{text
 | 
| 353 | "\<ASSUMES> \<exists>x. B x \<OBTAINS> x \<WHERE> B x"}'' in Isar. In the | |
| 354 | subsequent example we illustrate quantifier reasoning involving all | |
| 355 | four rules: | |
| 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 | |
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 358 | theorem | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 359 | assumes "\<exists>x. \<forall>y. R x y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 360 | shows "\<forall>y. \<exists>x. R x y" | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 361 | proof    -- {* @{text "\<forall>"} introduction *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 362 |   obtain x where "\<forall>y. R x y" using `\<exists>x. \<forall>y. R x y` ..    -- {* @{text "\<exists>"} elimination *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 363 |   fix y have "R x y" using `\<forall>y. R x y` ..    -- {* @{text "\<forall>"} destruction *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 364 |   then show "\<exists>x. R x y" ..    -- {* @{text "\<exists>"} introduction *}
 | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 365 | qed | 
| 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 366 | |
| 29734 | 367 | |
| 368 | subsection {* Canonical reasoning patterns *}
 | |
| 369 | ||
| 370 | text {*
 | |
| 371 | The main rules of first-order predicate logic from | |
| 372 |   \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant}
 | |
| 373 | can now be summarized as follows, using the native Isar statement | |
| 374 |   format of \secref{sec:framework-stmt}.
 | |
| 375 | ||
| 376 | \medskip | |
| 377 |   \begin{tabular}{l}
 | |
| 378 |   @{text "impI: \<ASSUMES> A \<Longrightarrow> B \<SHOWS> A \<longrightarrow> B"} \\
 | |
| 379 |   @{text "impD: \<ASSUMES> A \<longrightarrow> B \<AND> A \<SHOWS> B"} \\[1ex]
 | |
| 380 | ||
| 381 |   @{text "disjI\<^isub>1: \<ASSUMES> A \<SHOWS> A \<or> B"} \\
 | |
| 382 |   @{text "disjI\<^isub>2: \<ASSUMES> B \<SHOWS> A \<or> B"} \\
 | |
| 383 |   @{text "disjE: \<ASSUMES> A \<or> B \<OBTAINS> A \<BBAR> B"} \\[1ex]
 | |
| 384 | ||
| 385 |   @{text "conjI: \<ASSUMES> A \<AND> B \<SHOWS> A \<and> B"} \\
 | |
| 386 |   @{text "conjE: \<ASSUMES> A \<and> B \<OBTAINS> A \<AND> B"} \\[1ex]
 | |
| 387 | ||
| 388 |   @{text "falseE: \<ASSUMES> \<bottom> \<SHOWS> A"} \\
 | |
| 389 |   @{text "trueI: \<SHOWS> \<top>"} \\[1ex]
 | |
| 390 | ||
| 391 |   @{text "notI: \<ASSUMES> A \<Longrightarrow> \<bottom> \<SHOWS> \<not> A"} \\
 | |
| 392 |   @{text "notE: \<ASSUMES> \<not> A \<AND> A \<SHOWS> B"} \\[1ex]
 | |
| 393 | ||
| 394 |   @{text "allI: \<ASSUMES> \<And>x. B x \<SHOWS> \<forall>x. B x"} \\
 | |
| 395 |   @{text "allE: \<ASSUMES> \<forall>x. B x \<SHOWS> B a"} \\[1ex]
 | |
| 396 | ||
| 397 |   @{text "exI: \<ASSUMES> B a \<SHOWS> \<exists>x. B x"} \\
 | |
| 398 |   @{text "exE: \<ASSUMES> \<exists>x. B x \<OBTAINS> a \<WHERE> B a"}
 | |
| 399 |   \end{tabular}
 | |
| 400 | \medskip | |
| 401 | ||
| 402 | \noindent This essentially provides a declarative reading of Pure | |
| 403 | rules as Isar reasoning patterns: the rule statements tells how a | |
| 404 | canonical proof outline shall look like. Since the above rules have | |
| 29735 | 405 |   already been declared as @{attribute (Pure) intro}, @{attribute
 | 
| 406 |   (Pure) elim}, @{attribute (Pure) dest} --- each according to its
 | |
| 407 | particular shape --- we can immediately write Isar proof texts as | |
| 408 | follows: | |
| 29734 | 409 | *} | 
| 410 | ||
| 411 | (*<*) | |
| 412 | theorem "\<And>A. PROP A \<Longrightarrow> PROP A" | |
| 413 | proof - | |
| 414 | (*>*) | |
| 415 | ||
| 416 |   txt_raw {*\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 417 | ||
| 418 | have "A \<longrightarrow> B" | |
| 419 | proof | |
| 420 | assume A | |
| 421 | show B sorry %noproof | |
| 422 | qed | |
| 423 | ||
| 424 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 425 | ||
| 426 | have "A \<longrightarrow> B" and A sorry %noproof | |
| 427 | then have B .. | |
| 428 | ||
| 429 |   txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 430 | ||
| 431 | have A sorry %noproof | |
| 432 | then have "A \<or> B" .. | |
| 433 | ||
| 434 | have B sorry %noproof | |
| 435 | then have "A \<or> B" .. | |
| 436 | ||
| 437 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 438 | ||
| 439 | have "A \<or> B" sorry %noproof | |
| 440 | then have C | |
| 441 | proof | |
| 442 | assume A | |
| 443 | then show C sorry %noproof | |
| 444 | next | |
| 445 | assume B | |
| 446 | then show C sorry %noproof | |
| 447 | qed | |
| 448 | ||
| 449 |   txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 450 | ||
| 451 | have A and B sorry %noproof | |
| 452 | then have "A \<and> B" .. | |
| 453 | ||
| 454 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 455 | ||
| 456 | have "A \<and> B" sorry %noproof | |
| 457 | then obtain A and B .. | |
| 458 | ||
| 459 |   txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 460 | ||
| 461 | have "\<bottom>" sorry %noproof | |
| 462 | then have A .. | |
| 463 | ||
| 464 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 465 | ||
| 466 | have "\<top>" .. | |
| 467 | ||
| 468 |   txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 469 | ||
| 470 | have "\<not> A" | |
| 471 | proof | |
| 472 | assume A | |
| 473 | then show "\<bottom>" sorry %noproof | |
| 474 | qed | |
| 475 | ||
| 476 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 477 | ||
| 478 | have "\<not> A" and A sorry %noproof | |
| 479 | then have B .. | |
| 480 | ||
| 481 |   txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 482 | ||
| 483 | have "\<forall>x. B x" | |
| 484 | proof | |
| 485 | fix x | |
| 486 | show "B x" sorry %noproof | |
| 487 | qed | |
| 488 | ||
| 489 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 490 | ||
| 491 | have "\<forall>x. B x" sorry %noproof | |
| 492 | then have "B a" .. | |
| 493 | ||
| 494 |   txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 495 | ||
| 496 | have "\<exists>x. B x" | |
| 497 | proof | |
| 498 | show "B a" sorry %noproof | |
| 499 | qed | |
| 500 | ||
| 501 |   txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
 | |
| 502 | ||
| 503 | have "\<exists>x. B x" sorry %noproof | |
| 504 | then obtain a where "B a" .. | |
| 505 | ||
| 506 |   txt_raw {*\end{minipage}*}
 | |
| 507 | ||
| 508 | (*<*) | |
| 509 | qed | |
| 510 | (*>*) | |
| 511 | ||
| 512 | text {*
 | |
| 513 | \bigskip\noindent Of course, these proofs are merely examples. As | |
| 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. | |
| 518 | *} | |
| 519 | ||
| 29730 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 wenzelm parents: diff
changeset | 520 | end %visible |