# HG changeset patch # User wenzelm # Date 1167413145 -3600 # Node ID 9b772ac66830ae1eb6746e4579213e9d8661c7dd # Parent e5c96bb582525a0f65cca09fcc6b8fb3508eaa78 tuned specifications/proofs; diff -r e5c96bb58252 -r 9b772ac66830 src/FOL/ex/First_Order_Logic.thy --- a/src/FOL/ex/First_Order_Logic.thy Fri Dec 29 17:24:49 2006 +0100 +++ b/src/FOL/ex/First_Order_Logic.thy Fri Dec 29 18:25:45 2006 +0100 @@ -25,43 +25,43 @@ subsection {* Propositional logic *} -consts - false :: o ("\") - imp :: "o \ o \ o" (infixr "\" 25) - conj :: "o \ o \ o" (infixr "\" 35) - disj :: "o \ o \ o" (infixr "\" 30) - -axioms - falseE [elim]: "\ \ A" +axiomatization + false :: o ("\") and + imp :: "o \ o \ o" (infixr "\" 25) and + conj :: "o \ o \ o" (infixr "\" 35) and + disj :: "o \ o \ o" (infixr "\" 30) +where + falseE [elim]: "\ \ A" and - impI [intro]: "(A \ B) \ A \ B" - mp [dest]: "A \ B \ A \ B" + impI [intro]: "(A \ B) \ A \ B" and + mp [dest]: "A \ B \ A \ B" and - conjI [intro]: "A \ B \ A \ B" - conjD1: "A \ B \ A" - conjD2: "A \ B \ B" + conjI [intro]: "A \ B \ A \ B" and + conjD1: "A \ B \ A" and + conjD2: "A \ B \ B" and - disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" - disjI1 [intro]: "A \ A \ B" + disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" and + disjI1 [intro]: "A \ A \ B" and disjI2 [intro]: "B \ A \ B" -theorem conjE [elim]: "A \ B \ (A \ B \ C) \ C" -proof - - assume ab: "A \ B" - assume r: "A \ B \ C" - show C - proof (rule r) - from ab show A by (rule conjD1) - from ab show B by (rule conjD2) - qed +theorem conjE [elim]: + assumes "A \ B" + obtains A and B +proof + from `A \ B` show A by (rule conjD1) + from `A \ B` show B by (rule conjD2) qed -constdefs - true :: o ("\") +definition + true :: o ("\") where "\ \ \ \ \" - not :: "o \ o" ("\ _" [40] 40) + +definition + not :: "o \ o" ("\ _" [40] 40) where "\ A \ A \ \" - iff :: "o \ o \ o" (infixr "\" 25) + +definition + iff :: "o \ o \ o" (infixr "\" 25) where "A \ B \ (A \ B) \ (B \ A)" @@ -73,44 +73,43 @@ theorem notI [intro]: "(A \ \) \ \ A" proof (unfold not_def) assume "A \ \" - thus "A \ \" .. + then show "A \ \" .. qed theorem notE [elim]: "\ A \ A \ B" proof (unfold not_def) assume "A \ \" and A - hence \ .. thus B .. + then have \ .. then show B .. qed theorem iffI [intro]: "(A \ B) \ (B \ A) \ A \ B" proof (unfold iff_def) - assume "A \ B" hence "A \ B" .. - moreover assume "B \ A" hence "B \ A" .. + assume "A \ B" then have "A \ B" .. + moreover assume "B \ A" then have "B \ A" .. ultimately show "(A \ B) \ (B \ A)" .. qed theorem iff1 [elim]: "A \ B \ A \ B" proof (unfold iff_def) assume "(A \ B) \ (B \ A)" - hence "A \ B" .. - thus "A \ B" .. + then have "A \ B" .. + then show "A \ B" .. qed theorem iff2 [elim]: "A \ B \ B \ A" proof (unfold iff_def) assume "(A \ B) \ (B \ A)" - hence "B \ A" .. - thus "B \ A" .. + then have "B \ A" .. + then show "B \ A" .. qed subsection {* Equality *} -consts - equal :: "i \ i \ o" (infixl "=" 50) - -axioms - refl [intro]: "x = x" +axiomatization + equal :: "i \ i \ o" (infixl "=" 50) +where + refl [intro]: "x = x" and subst: "x = y \ P(x) \ P(y)" theorem trans [trans]: "x = y \ y = z \ x = z" @@ -125,38 +124,36 @@ subsection {* Quantifiers *} -consts - All :: "(i \ o) \ o" (binder "\" 10) - Ex :: "(i \ o) \ o" (binder "\" 10) - -axioms - allI [intro]: "(\x. P(x)) \ \x. P(x)" - allD [dest]: "\x. P(x) \ P(a)" - - exI [intro]: "P(a) \ \x. P(x)" +axiomatization + All :: "(i \ o) \ o" (binder "\" 10) and + Ex :: "(i \ o) \ o" (binder "\" 10) +where + allI [intro]: "(\x. P(x)) \ \x. P(x)" and + allD [dest]: "\x. P(x) \ P(a)" and + exI [intro]: "P(a) \ \x. P(x)" and exE [elim]: "\x. P(x) \ (\x. P(x) \ C) \ C" lemma "(\x. P(f(x))) \ (\y. P(y))" proof assume "\x. P(f(x))" - thus "\y. P(y)" + then show "\y. P(y)" proof fix x assume "P(f(x))" - thus ?thesis .. + then show ?thesis .. qed qed lemma "(\x. \y. R(x, y)) \ (\y. \x. R(x, y))" proof assume "\x. \y. R(x, y)" - thus "\y. \x. R(x, y)" + then show "\y. \x. R(x, y)" proof fix x assume a: "\y. R(x, y)" show ?thesis proof fix y from a have "R(x, y)" .. - thus "\x. R(x, y)" .. + then show "\x. R(x, y)" .. qed qed qed