| author | wenzelm | 
| Fri, 25 Sep 2020 11:24:28 +0200 | |
| changeset 72293 | 584aea0b29bb | 
| parent 71924 | e5df9c8d9d4b | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 71924 
e5df9c8d9d4b
clarified sessions: "Notable Examples in Isabelle/Pure";
 wenzelm parents: 
63585diff
changeset | 1 | (* Title: Pure/Examples/First_Order_Logic.thy | 
| 61758 | 2 | Author: Makarius | 
| 12369 | 3 | *) | 
| 4 | ||
| 60770 | 5 | section \<open>A simple formulation of First-Order Logic\<close> | 
| 12369 | 6 | |
| 60770 | 7 | text \<open> | 
| 61758 | 8 | The subsequent theory development illustrates single-sorted intuitionistic | 
| 61935 | 9 | first-order logic with equality, formulated within the Pure framework. | 
| 60770 | 10 | \<close> | 
| 12369 | 11 | |
| 61758 | 12 | theory First_Order_Logic | 
| 63585 | 13 | imports Pure | 
| 61758 | 14 | begin | 
| 15 | ||
| 16 | subsection \<open>Abstract syntax\<close> | |
| 12369 | 17 | |
| 18 | typedecl i | |
| 19 | typedecl o | |
| 20 | ||
| 61758 | 21 | judgment Trueprop :: "o \<Rightarrow> prop"  ("_" 5)
 | 
| 12369 | 22 | |
| 23 | ||
| 60770 | 24 | subsection \<open>Propositional logic\<close> | 
| 12369 | 25 | |
| 61758 | 26 | axiomatization false :: o  ("\<bottom>")
 | 
| 27 | where falseE [elim]: "\<bottom> \<Longrightarrow> A" | |
| 28 | ||
| 12369 | 29 | |
| 61758 | 30 | axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25) | 
| 31 | where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B" | |
| 32 | and mp [dest]: "A \<longrightarrow> B \<Longrightarrow> A \<Longrightarrow> B" | |
| 33 | ||
| 12369 | 34 | |
| 61758 | 35 | axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) | 
| 36 | where conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B" | |
| 37 | and conjD1: "A \<and> B \<Longrightarrow> A" | |
| 38 | and conjD2: "A \<and> B \<Longrightarrow> B" | |
| 12369 | 39 | |
| 21939 | 40 | theorem conjE [elim]: | 
| 41 | assumes "A \<and> B" | |
| 42 | obtains A and B | |
| 43 | proof | |
| 61758 | 44 | from \<open>A \<and> B\<close> show A | 
| 45 | by (rule conjD1) | |
| 46 | from \<open>A \<and> B\<close> show B | |
| 47 | by (rule conjD2) | |
| 12369 | 48 | qed | 
| 49 | ||
| 61758 | 50 | |
| 51 | axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) | |
| 52 | where disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C" | |
| 53 | and disjI1 [intro]: "A \<Longrightarrow> A \<or> B" | |
| 54 | and disjI2 [intro]: "B \<Longrightarrow> A \<or> B" | |
| 55 | ||
| 56 | ||
| 60769 | 57 | definition true :: o  ("\<top>")
 | 
| 58 | where "\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>" | |
| 21939 | 59 | |
| 61758 | 60 | theorem trueI [intro]: \<top> | 
| 61 | unfolding true_def .. | |
| 62 | ||
| 63 | ||
| 60769 | 64 | definition not :: "o \<Rightarrow> o"  ("\<not> _" [40] 40)
 | 
| 65 | where "\<not> A \<equiv> A \<longrightarrow> \<bottom>" | |
| 21939 | 66 | |
| 12369 | 67 | theorem notI [intro]: "(A \<Longrightarrow> \<bottom>) \<Longrightarrow> \<not> A" | 
| 61758 | 68 | unfolding not_def .. | 
| 12369 | 69 | |
| 70 | theorem notE [elim]: "\<not> A \<Longrightarrow> A \<Longrightarrow> B" | |
| 61758 | 71 | unfolding not_def | 
| 72 | proof - | |
| 12369 | 73 | assume "A \<longrightarrow> \<bottom>" and A | 
| 60769 | 74 | then have \<bottom> .. | 
| 75 | then show B .. | |
| 12369 | 76 | qed | 
| 77 | ||
| 61758 | 78 | |
| 79 | definition iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longleftrightarrow>" 25) | |
| 80 | where "A \<longleftrightarrow> B \<equiv> (A \<longrightarrow> B) \<and> (B \<longrightarrow> A)" | |
| 81 | ||
| 82 | theorem iffI [intro]: | |
| 83 | assumes "A \<Longrightarrow> B" | |
| 84 | and "B \<Longrightarrow> A" | |
| 85 | shows "A \<longleftrightarrow> B" | |
| 86 | unfolding iff_def | |
| 87 | proof | |
| 88 | from \<open>A \<Longrightarrow> B\<close> show "A \<longrightarrow> B" .. | |
| 89 | from \<open>B \<Longrightarrow> A\<close> show "B \<longrightarrow> A" .. | |
| 12392 | 90 | qed | 
| 91 | ||
| 61758 | 92 | theorem iff1 [elim]: | 
| 93 | assumes "A \<longleftrightarrow> B" and A | |
| 94 | shows B | |
| 95 | proof - | |
| 96 | from \<open>A \<longleftrightarrow> B\<close> have "(A \<longrightarrow> B) \<and> (B \<longrightarrow> A)" | |
| 97 | unfolding iff_def . | |
| 21939 | 98 | then have "A \<longrightarrow> B" .. | 
| 61758 | 99 | from this and \<open>A\<close> show B .. | 
| 12392 | 100 | qed | 
| 101 | ||
| 61758 | 102 | theorem iff2 [elim]: | 
| 103 | assumes "A \<longleftrightarrow> B" and B | |
| 104 | shows A | |
| 105 | proof - | |
| 106 | from \<open>A \<longleftrightarrow> B\<close> have "(A \<longrightarrow> B) \<and> (B \<longrightarrow> A)" | |
| 107 | unfolding iff_def . | |
| 21939 | 108 | then have "B \<longrightarrow> A" .. | 
| 61758 | 109 | from this and \<open>B\<close> show A .. | 
| 12392 | 110 | qed | 
| 111 | ||
| 12369 | 112 | |
| 60770 | 113 | subsection \<open>Equality\<close> | 
| 12369 | 114 | |
| 61758 | 115 | axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infixl "=" 50) | 
| 116 | where refl [intro]: "x = x" | |
| 117 | and subst: "x = y \<Longrightarrow> P x \<Longrightarrow> P y" | |
| 12369 | 118 | |
| 119 | theorem trans [trans]: "x = y \<Longrightarrow> y = z \<Longrightarrow> x = z" | |
| 120 | by (rule subst) | |
| 121 | ||
| 122 | theorem sym [sym]: "x = y \<Longrightarrow> y = x" | |
| 123 | proof - | |
| 124 | assume "x = y" | |
| 61758 | 125 | from this and refl show "y = x" | 
| 126 | by (rule subst) | |
| 12369 | 127 | qed | 
| 128 | ||
| 129 | ||
| 60770 | 130 | subsection \<open>Quantifiers\<close> | 
| 12369 | 131 | |
| 61758 | 132 | axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10) | 
| 133 | where allI [intro]: "(\<And>x. P x) \<Longrightarrow> \<forall>x. P x" | |
| 134 | and allD [dest]: "\<forall>x. P x \<Longrightarrow> P a" | |
| 135 | ||
| 136 | axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10) | |
| 137 | where exI [intro]: "P a \<Longrightarrow> \<exists>x. P x" | |
| 138 | and exE [elim]: "\<exists>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> C) \<Longrightarrow> C" | |
| 12369 | 139 | |
| 140 | ||
| 26958 | 141 | lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)" | 
| 12369 | 142 | proof | 
| 26958 | 143 | assume "\<exists>x. P (f x)" | 
| 61758 | 144 | then obtain x where "P (f x)" .. | 
| 145 | then show "\<exists>y. P y" .. | |
| 12369 | 146 | qed | 
| 147 | ||
| 26958 | 148 | lemma "(\<exists>x. \<forall>y. R x y) \<longrightarrow> (\<forall>y. \<exists>x. R x y)" | 
| 12369 | 149 | proof | 
| 26958 | 150 | assume "\<exists>x. \<forall>y. R x y" | 
| 61758 | 151 | then obtain x where "\<forall>y. R x y" .. | 
| 152 | show "\<forall>y. \<exists>x. R x y" | |
| 12369 | 153 | proof | 
| 61758 | 154 | fix y | 
| 155 | from \<open>\<forall>y. R x y\<close> have "R x y" .. | |
| 156 | then show "\<exists>x. R x y" .. | |
| 12369 | 157 | qed | 
| 158 | qed | |
| 159 | ||
| 160 | end |