# HG changeset patch # User paulson # Date 1058785327 -7200 # Node ID 3a73850c6c7d4f439932d23ff3e97a020c74b9e6 # Parent fb9c392644a15e3ca05a5941d1b4b825785cb167 Tidied some examples diff -r fb9c392644a1 -r 3a73850c6c7d src/ZF/OrderArith.thy --- a/src/ZF/OrderArith.thy Mon Jul 21 10:58:16 2003 +0200 +++ b/src/ZF/OrderArith.thy Mon Jul 21 13:02:07 2003 +0200 @@ -543,6 +543,27 @@ apply (frule ok, assumption+, blast) done +subsubsection{*Bijections involving Powersets*} + +lemma Pow_sum_bij: + "(\Z \ Pow(A+B). <{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}>) + \ bij(Pow(A+B), Pow(A)*Pow(B))" +apply (rule_tac d = "%. {Inl (x). x \ X} Un {Inr (y). y \ Y}" + in lam_bijective) +apply force+ +done + +text{*As a special case, we have @{term "bij(Pow(A*B), A -> Pow(B))"} *} +lemma Pow_Sigma_bij: + "(\r \ Pow(Sigma(A,B)). \x \ A. r``{x}) + \ bij(Pow(Sigma(A,B)), \x \ A. Pow(B(x)))" +apply (rule_tac d = "%f. \x \ A. \y \ f`x. {}" in lam_bijective) +apply (blast intro: lam_type) +apply (blast dest: apply_type, simp_all) +apply fast (*strange, but blast can't do it*) +apply (rule fun_extension, auto) +by blast + ML {* val measure_def = thm "measure_def"; diff -r fb9c392644a1 -r 3a73850c6c7d src/ZF/ex/misc.thy --- a/src/ZF/ex/misc.thy Mon Jul 21 10:58:16 2003 +0200 +++ b/src/ZF/ex/misc.thy Mon Jul 21 13:02:07 2003 +0200 @@ -3,32 +3,37 @@ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge -Miscellaneous examples for Zermelo-Fraenkel Set Theory Composition of homomorphisms, Pastre's examples, ... *) +header{*Miscellaneous ZF Examples*} + theory misc = Main: - +subsection{*Various Small Problems*} -(*These two are cited in Benzmueller and Kohlhase's system description of LEO, - CADE-15, 1998 (page 139-143) as theorems LEO could not prove.*) +text{*A weird property of ordered pairs.*} +lemma "b\c ==> Int = " +by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast) + +text{*These two are cited in Benzmueller and Kohlhase's system description of + LEO, CADE-15, 1998 (page 139-143) as theorems LEO could not prove.*} lemma "(X = Y Un Z) <-> (Y \ X & Z \ X & (\V. Y \ V & Z \ V --> X \ V))" by (blast intro!: equalityI) -(*the dual of the previous one*) +text{*the dual of the previous one} lemma "(X = Y Int Z) <-> (X \ Y & X \ Z & (\V. V \ Y & V \ Z --> V \ X))" by (blast intro!: equalityI) -(*trivial example of term synthesis: apparently hard for some provers!*) +text{*trivial example of term synthesis: apparently hard for some provers!} lemma "a \ b ==> a:?X & b \ ?X" by blast -(*Nice Blast_tac benchmark. Proved in 0.3s; old tactics can't manage it!*) +text{*Nice Blast_tac benchmark. Proved in 0.3s; old tactics can't manage it!} lemma "\x \ S. \y \ S. x \ y ==> \z. S \ {z}" by blast -(*variant of the benchmark above*) +text{*variant of the benchmark above} lemma "\x \ S. Union(S) \ x ==> \z. S \ {z}" by blast @@ -39,16 +44,19 @@ lemma "(\F. {x} \ F --> {y} \ F) --> (\A. x \ A --> y \ A)" by best +text{*A characterization of functions suggested by Tobias Nipkow*} +lemma "r \ domain(r)->B <-> r \ domain(r)*B & (\X. r `` (r -`` X) \ X)" +by (unfold Pi_def function_def, best) -(*** Composition of homomorphisms is a homomorphism ***) -(*Given as a challenge problem in +subsection{*Composition of homomorphisms is a Homomorphism*} + +text{*Given as a challenge problem in R. Boyer et al., Set Theory in First-Order Logic: Clauses for G\"odel's Axioms, - JAR 2 (1986), 287-327 -*) + JAR 2 (1986), 287-327 *} -(*collecting the relevant lemmas*) +text{*collecting the relevant lemmas} declare comp_fun [simp] SigmaI [simp] apply_funtype [simp] (*Force helps prove conditions of rewrites such as comp_fun_apply, since @@ -60,7 +68,7 @@ (K O J) \ hom(A,f,C,h)" by force -(*Another version, with meta-level rewriting*) +text{*Another version, with meta-level rewriting} lemma "(!! A f B g. hom(A,f,B,g) == {H \ A->B. f \ A*A->A & g \ B*B->B & (\x \ A. \y \ A. H`(f`) = g`)}) @@ -68,17 +76,11 @@ by force - -(** A characterization of functions suggested by Tobias Nipkow **) - -lemma "r \ domain(r)->B <-> r \ domain(r)*B & (\X. r `` (r -`` X) \ X)" -by (unfold Pi_def function_def, best) +subsection{*Pastre's Examples*} -(**** From D Pastre. Automatic theorem proving in set theory. - Artificial Intelligence, 10:1--27, 1978. - - Previously, these were done using ML code, but blast manages fine. -****) +text{*D Pastre. Automatic theorem proving in set theory. + Artificial Intelligence, 10:1--27, 1978. +Previously, these were done using ML code, but blast manages fine.*} lemmas compIs [intro] = comp_surj comp_inj comp_fun [intro] lemmas compDs [dest] = comp_mem_injD1 comp_mem_surjD1 @@ -120,26 +122,5 @@ by (unfold bij_def, blast) -(** Yet another example... **) - -lemma Pow_sum_bij: - "(\Z \ Pow(A+B). <{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}>) - \ bij(Pow(A+B), Pow(A)*Pow(B))" -apply (rule_tac d = "%. {Inl (x). x \ X} Un {Inr (y). y \ Y}" - in lam_bijective) -apply force+ -done - -(*As a special case, we have bij(Pow(A*B), A -> Pow B) *) -lemma Pow_Sigma_bij: - "(\r \ Pow(Sigma(A,B)). \x \ A. r``{x}) - \ bij(Pow(Sigma(A,B)), \x \ A. Pow(B(x)))" -apply (rule_tac d = "%f. \x \ A. \y \ f`x. {}" in lam_bijective) -apply (blast intro: lam_type) -apply (blast dest: apply_type, simp_all) -apply fast (*strange, but blast can't do it*) -apply (rule fun_extension, auto) -by blast - end