theory Demo4 = Main: text{* A warming up exercise: propositional logic. *} lemma "(A \ (B \ C)) \ (A \ B \ C)" oops -- {* more rules *} text{* = (iff) *} thm iffI iffE iffD1 iffD2 text{* Equality *} thm refl sym trans text{* \ *} thm notI notE text{* True and False *} thm TrueI FalseE -- ----------------------------- text {* defer and prefer *} lemma "(A \ A) = (A \ A)" apply (rule iffI) defer sorry lemma contraposition: "(A \ B) \ (\B \ \A)" sorry text {* classical, cases *} lemma "(\A \ B) \ (\B \ A)" sorry lemma Pierce: "((A \ B) \ A) \ A" sorry lemma PnP: "P \ \P" apply (cases "P") oops text {* safe and unsafe *} lemma "A \ C \ A \ B" oops -- ------------------------------------ text {* Quantifier reasoning *} text{* A successful proof: *} lemma "\x. \y. x = y" apply(rule allI) apply(rule exI) apply(rule refl) done text{* An unsuccessful proof: *} lemma "\y. \x. x = y" apply(rule exI) apply(rule allI) (* Does not work: apply(rule refl) *) oops text{* Intro and elim resoning: *} lemma "\y. \x. P x y \ \x. \y. P x y" (* the safe rules first: *) apply(rule allI) apply(erule exE) (* now the unsafe ones: *) apply(rule_tac x=y in exI) apply(erule_tac x=x in allE) apply(assumption) done text{* What happens if an unsafe rule is tried too early: *} lemma "\y. \x. P x y \ \x. \y. P x y" apply(rule allI) apply(rule exI) apply(erule exE) apply(erule allE) (* Fails now: apply(assumption) *) oops text {* Instantiation in more detail: *} text{* Instantiating allE: *} lemma "\x. P x \ P 37" thm allE apply (erule_tac x = "37" in allE) apply assumption done text{* Instantiating exI: *} lemma "\n. P (f n) \ \m. P m" apply(erule exE) thm exI apply(rule_tac x = "f n" in exI) apply assumption done text{* Instantiation removes ambiguity: *} lemma "\ A \ B; C \ D \ \ D" thm conjE apply(erule_tac P = "C" in conjE) (* without instantiation, the wrong one is chosen (first) *) apply assumption done text {* Instantiation with "where" and "of" *} thm conjI thm conjI [of A B] thm conjI [where Q = "f x"] text {* Exercises *} lemma "\x. \y. P x y \ \y. \x. P x y" oops lemma "(\x. P x) \ Q \ \x. P x \ Q" oops lemma "\x. (P x \ (\x. P x))" oops -- ---------------------------------------------- text{* Renaming parameters: *} lemma "\x y z. P x y z" apply(rename_tac a b) oops lemma "\x. P x \ \x. \x. P x" apply(rule allI) apply(rule allI) apply(rename_tac X) apply(erule_tac x=X in allE) apply assumption done text {* Forward reasoning: drule/frule/OF/THEN*} lemma "A \ B \ \ \ A" thm conjunct1 apply (drule conjunct1) apply (rule notI) apply (erule notE) apply assumption done lemma "\x. P x \ P t" thm spec apply (frule spec) apply assumption done thm dvd_add dvd_refl thm dvd_add [OF dvd_refl] thm dvd_add [OF dvd_refl dvd_refl] -- --------------------------------------------- text {* Epsilon *} lemma "(\x. P x) = P (SOME x. P x)" apply (rule iffI) apply (erule exE) apply (rule someI) apply assumption apply (rule exI) apply assumption done text {* Automation *} lemma "\x y. P x y \ Q x y \ R x y" apply (intro allI conjI) oops lemma "\x y. P x y \ Q x y \ R x y" apply clarify oops lemma "\x y. P x y \ Q x y \ R x y" apply safe oops lemma "\y. \x. P x y \ \x. \y. P x y" apply blast done lemma "\y. \x. P x y \ \x. \y. P x y" apply fast done end