diff -r 000000000000 -r 7949f97df77a ex/set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/set.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,125 @@ +(* Title: HOL/ex/set.ML + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Cantor's Theorem; the Schroeder-Berstein Theorem. +*) + + +writeln"File HOL/ex/set."; + +(*** Cantor's Theorem: There is no surjection from a set to its powerset. ***) + +goal Set.thy "~ (? f:: 'a=>'a set. ! S. ? x. f(x) = S)"; +(*requires best-first search because it is undirectional*) +by (best_tac (set_cs addSEs [equalityCE]) 1); +val cantor1 = result(); + +(*This form displays the diagonal term*) +goal Set.thy "! f:: 'a=>'a set. ! x. ~ f(x) = ?S(f)"; +by (best_tac (set_cs addSEs [equalityCE]) 1); +uresult(); + +(*This form exploits the set constructs*) +goal Set.thy "~ ?S : range(f :: 'a=>'a set)"; +by (rtac notI 1); +by (etac rangeE 1); +by (etac equalityCE 1); +by (dtac CollectD 1); +by (contr_tac 1); +by (swap_res_tac [CollectI] 1); +by (assume_tac 1); + +choplev 0; +by (best_tac (set_cs addSEs [equalityCE]) 1); + +(*** The Schroder-Berstein Theorem ***) + +val prems = goalw Lfp.thy [image_def] "inj(f) ==> Inv(f)``(f``X) = X"; +by (cut_facts_tac prems 1); +by (rtac equalityI 1); +by (fast_tac (set_cs addEs [Inv_f_f RS ssubst]) 1); +by (fast_tac (set_cs addEs [Inv_f_f RS ssubst]) 1); +val inv_image_comp = result(); + +val prems = goal Set.thy "~ f(a) : (f``X) ==> ~ a:X"; +by (cfast_tac prems 1); +val contra_imageI = result(); + +goal Lfp.thy "(~ a : Compl(X)) = (a:X)"; +by (fast_tac set_cs 1); +val not_Compl = result(); + +(*Lots of backtracking in this proof...*) +val [compl,fg,Xa] = goal Lfp.thy + "[| Compl(f``X) = g``Compl(X); f(a)=g(b); a:X |] ==> b:X"; +by (EVERY1 [rtac (not_Compl RS subst), rtac contra_imageI, + rtac (compl RS subst), rtac (fg RS subst), stac not_Compl, + rtac imageI, rtac Xa]); +val disj_lemma = result(); + +goal Lfp.thy "range(%z. if(z:X, f(z), g(z))) = f``X Un g``Compl(X)"; +by (rtac equalityI 1); +by (rewtac range_def); +by (fast_tac (set_cs addIs [if_P RS sym, if_not_P RS sym]) 2); +by (rtac subsetI 1); +by (etac CollectE 1); +by (etac exE 1); +by (etac ssubst 1); +by (rtac (excluded_middle RS disjE) 1); +by (EVERY' [rtac (if_P RS ssubst), atac, fast_tac set_cs] 2); +by (EVERY' [rtac (if_not_P RS ssubst), atac, fast_tac set_cs] 1); +val range_if_then_else = result(); + +goal Lfp.thy "a : X Un Compl(X)"; +by (fast_tac set_cs 1); +val X_Un_Compl = result(); + +goalw Lfp.thy [surj_def] "surj(f) = (!a. a : range(f))"; +by (fast_tac (set_cs addEs [ssubst]) 1); +val surj_iff_full_range = result(); + +val [compl] = goal Lfp.thy + "Compl(f``X) = g``Compl(X) ==> surj(%z. if(z:X, f(z), g(z)))"; +by (sstac [surj_iff_full_range, range_if_then_else, compl RS sym] 1); +by (rtac (X_Un_Compl RS allI) 1); +val surj_if_then_else = result(); + +val [injf,injg,compl,bij] = goal Lfp.thy + "[| inj_onto(f,X); inj_onto(g,Compl(X)); Compl(f``X) = g``Compl(X); \ +\ bij = (%z. if(z:X, f(z), g(z))) |] ==> \ +\ inj(bij) & surj(bij)"; +val f_eq_gE = make_elim (compl RS disj_lemma); +by (rtac (bij RS ssubst) 1); +by (rtac conjI 1); +by (rtac (compl RS surj_if_then_else) 2); +by (rewtac inj_def); +by (cut_facts_tac [injf,injg] 1); +by (EVERY1 [rtac allI, rtac allI, stac expand_if, rtac conjI, stac expand_if]); +by (fast_tac (set_cs addEs [inj_ontoD, sym RS f_eq_gE]) 1); +by (stac expand_if 1); +by (fast_tac (set_cs addEs [inj_ontoD, f_eq_gE]) 1); +val bij_if_then_else = result(); + +goal Lfp.thy "? X. X = Compl(g``Compl(f:: 'a=>'b``X))"; +by (rtac exI 1); +by (rtac lfp_Tarski 1); +by (REPEAT (ares_tac [monoI, image_mono, Compl_anti_mono] 1)); +val decomposition = result(); + +val [injf,injg] = goal Lfp.thy + "[| inj(f:: 'a=>'b); inj(g:: 'b=>'a) |] ==> \ +\ ? h:: 'a=>'b. inj(h) & surj(h)"; +by (rtac (decomposition RS exE) 1); +by (rtac exI 1); +by (rtac bij_if_then_else 1); +by (EVERY [rtac refl 4, rtac (injf RS inj_imp) 1, + rtac (injg RS inj_onto_Inv) 1]); +by (EVERY1 [etac ssubst, stac double_complement, rtac subsetI, + etac imageE, etac ssubst, rtac rangeI]); +by (EVERY1 [etac ssubst, stac double_complement, + rtac (injg RS inv_image_comp RS sym)]); +val schroeder_bernstein = result(); + +writeln"Reached end of file.";