diff -r 2ec9c808a8a7 -r 8eb12693cead doc-src/TutorialI/Rules/Blast.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Rules/Blast.thy Mon Oct 23 16:25:04 2000 +0200 @@ -0,0 +1,40 @@ +theory Blast = Main: + +lemma "((\x. \y. p(x)=p(y)) = ((\x. q(x))=(\y. p(y)))) = + ((\x. \y. q(x)=q(y)) = ((\x. p(x))=(\y. q(y))))" +apply blast +done + +text{*\noindent Until now, we have proved everything using only induction and +simplification. Substantial proofs require more elaborate types of +inference.*} + +lemma "(\x. honest(x) \ industrious(x) \ healthy(x)) \ + \ (\x. grocer(x) \ healthy(x)) \ + (\x. industrious(x) \ grocer(x) \ honest(x)) \ + (\x. cyclist(x) \ industrious(x)) \ + (\x. \healthy(x) \ cyclist(x) \ \honest(x)) + \ (\x. grocer(x) \ \cyclist(x))"; +apply blast +done + +lemma "(\i\I. A(i)) \ (\j\J. B(j)) = + (\i\I. \j\J. A(i) \ B(j))" +apply blast +done + +text {* +@{thm[display] mult_is_0} + \rulename{mult_is_0}} + +@{thm[display] finite_Un} + \rulename{finite_Un}} +*}; + + +lemma [iff]: "(xs@ys = []) = (xs=[] & ys=[])" + apply (induct_tac xs) + by (simp_all); + +(*ideas for uses of intro, etc.: ex/Primes/is_gcd_unique?*) +end