# HG changeset patch # User nipkow # Date 860164408 -7200 # Node ID 3fac3e8d5d3e97cfe816582f70418c31e25cda8e # Parent 8a680e310f041ec3024d86a4a599784e4f38d13e moved inj and surj from Set to Fun and Inv -> inv. diff -r 8a680e310f04 -r 3fac3e8d5d3e src/HOL/Fun.ML --- a/src/HOL/Fun.ML Fri Apr 04 16:27:39 1997 +0200 +++ b/src/HOL/Fun.ML Fri Apr 04 16:33:28 1997 +0200 @@ -19,48 +19,6 @@ qed "apply_inverse"; -(*** Image of a set under a function ***) - -(*Frequently b does not have the syntactic form of f(x).*) -val prems = goalw Fun.thy [image_def] "[| b=f(x); x:A |] ==> b : f``A"; -by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1)); -qed "image_eqI"; - -bind_thm ("imageI", refl RS image_eqI); - -(*The eta-expansion gives variable-name preservation.*) -val major::prems = goalw Fun.thy [image_def] - "[| b : (%x.f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; -by (rtac (major RS CollectD RS bexE) 1); -by (REPEAT (ares_tac prems 1)); -qed "imageE"; - -AddIs [image_eqI]; -AddSEs [imageE]; - -goalw Fun.thy [o_def] "(f o g)``r = f``(g``r)"; -by (Fast_tac 1); -qed "image_compose"; - -goal Fun.thy "f``(A Un B) = f``A Un f``B"; -by (Fast_tac 1); -qed "image_Un"; - -(*** Range of a function -- just a translation for image! ***) - -goal Fun.thy "!!b. b=f(x) ==> b : range(f)"; -by (EVERY1 [etac image_eqI, rtac UNIV_I]); -bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI)); - -bind_thm ("rangeI", UNIV_I RS imageI); - -val [major,minor] = goal Fun.thy - "[| b : range(%x.f(x)); !!x. b=f(x) ==> P |] ==> P"; -by (rtac (major RS imageE) 1); -by (etac minor 1); -qed "rangeE"; - - (*** inj(f): f is a one-to-one function ***) val prems = goalw Fun.thy [inj_def] @@ -95,14 +53,14 @@ qed "inj_select"; (*A one-to-one function has an inverse (given using select).*) -val [major] = goalw Fun.thy [Inv_def] "inj(f) ==> Inv f (f x) = x"; +val [major] = goalw Fun.thy [inv_def] "inj(f) ==> inv f (f x) = x"; by (EVERY1 [rtac (major RS inj_select)]); -qed "Inv_f_f"; +qed "inv_f_f"; (* Useful??? *) val [oneone,minor] = goal Fun.thy - "[| inj(f); !!y. y: range(f) ==> P(Inv f y) |] ==> P(x)"; -by (res_inst_tac [("t", "x")] (oneone RS (Inv_f_f RS subst)) 1); + "[| inj(f); !!y. y: range(f) ==> P(inv f y) |] ==> P(x)"; +by (res_inst_tac [("t", "x")] (oneone RS (inv_f_f RS subst)) 1); by (rtac (rangeI RS minor) 1); qed "inj_transfer"; @@ -152,27 +110,22 @@ by (fast_tac (!claset addIs [prem RS injD, inj_ontoI]) 1); qed "inj_imp"; -val [prem] = goalw Fun.thy [Inv_def] "y : range(f) ==> f(Inv f y) = y"; +val [prem] = goalw Fun.thy [inv_def] "y : range(f) ==> f(inv f y) = y"; by (EVERY1 [rtac (prem RS rangeE), rtac selectI, etac sym]); -qed "f_Inv_f"; - -val prems = goal Fun.thy - "[| Inv f x=Inv f y; x: range(f); y: range(f) |] ==> x=y"; -by (rtac (arg_cong RS box_equals) 1); -by (REPEAT (resolve_tac (prems @ [f_Inv_f]) 1)); -qed "Inv_injective"; +qed "f_inv_f"; val prems = goal Fun.thy - "[| inj(f); A<=range(f) |] ==> inj_onto (Inv f) A"; + "[| inv f x=inv f y; x: range(f); y: range(f) |] ==> x=y"; +by (rtac (arg_cong RS box_equals) 1); +by (REPEAT (resolve_tac (prems @ [f_inv_f]) 1)); +qed "inv_injective"; + +val prems = goal Fun.thy + "[| inj(f); A<=range(f) |] ==> inj_onto (inv f) A"; by (cut_facts_tac prems 1); by (fast_tac (!claset addIs [inj_ontoI] - addEs [Inv_injective,injD]) 1); -qed "inj_onto_Inv"; + addEs [inv_injective,injD]) 1); +qed "inj_onto_inv"; -AddIs [rangeI]; -AddSEs [rangeE]; - val set_cs = !claset delrules [equalityI]; - - diff -r 8a680e310f04 -r 3fac3e8d5d3e src/HOL/Fun.thy --- a/src/HOL/Fun.thy Fri Apr 04 16:27:39 1997 +0200 +++ b/src/HOL/Fun.thy Fri Apr 04 16:33:28 1997 +0200 @@ -3,7 +3,22 @@ Author: Tobias Nipkow, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge -Lemmas about functions. +Notions about functions. *) -Fun = Set +Fun = Set + + +consts + + inj, surj :: ('a => 'b) => bool (*inj/surjective*) + inj_onto :: ['a => 'b, 'a set] => bool + inv :: ('a => 'b) => ('b => 'a) + +defs + + inj_def "inj f == ! x y. f(x)=f(y) --> x=y" + inj_onto_def "inj_onto f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" + surj_def "surj f == ! y. ? x. y=f(x)" + inv_def "inv(f::'a=>'b) == (% y. @x. f(x)=y)" + +end diff -r 8a680e310f04 -r 3fac3e8d5d3e src/HOL/HOL.thy --- a/src/HOL/HOL.thy Fri Apr 04 16:27:39 1997 +0200 +++ b/src/HOL/HOL.thy Fri Apr 04 16:33:28 1997 +0200 @@ -33,7 +33,6 @@ Not :: bool => bool ("~ _" [40] 40) True, False :: bool If :: [bool, 'a, 'a] => 'a ("(if (_)/ then (_)/ else (_))" 10) - Inv :: ('a => 'b) => ('b => 'a) (* Binders *) @@ -170,7 +169,6 @@ (* Misc Definitions *) Let_def "Let s f == f(s)" - Inv_def "Inv(f::'a=>'b) == (% y. @x. f(x)=y)" o_def "(f::'b=>'c) o g == (%(x::'a). f(g(x)))" if_def "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)" diff -r 8a680e310f04 -r 3fac3e8d5d3e src/HOL/Set.ML --- a/src/HOL/Set.ML Fri Apr 04 16:27:39 1997 +0200 +++ b/src/HOL/Set.ML Fri Apr 04 16:33:28 1997 +0200 @@ -602,6 +602,51 @@ AddEs [InterD, InterE]; +(*** Image of a set under a function ***) + +(*Frequently b does not have the syntactic form of f(x).*) +val prems = goalw thy [image_def] "[| b=f(x); x:A |] ==> b : f``A"; +by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1)); +qed "image_eqI"; + +bind_thm ("imageI", refl RS image_eqI); + +(*The eta-expansion gives variable-name preservation.*) +val major::prems = goalw thy [image_def] + "[| b : (%x.f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; +by (rtac (major RS CollectD RS bexE) 1); +by (REPEAT (ares_tac prems 1)); +qed "imageE"; + +AddIs [image_eqI]; +AddSEs [imageE]; + +goalw thy [o_def] "(f o g)``r = f``(g``r)"; +by (Fast_tac 1); +qed "image_compose"; + +goal thy "f``(A Un B) = f``A Un f``B"; +by (Fast_tac 1); +qed "image_Un"; + + +(*** Range of a function -- just a translation for image! ***) + +goal thy "!!b. b=f(x) ==> b : range(f)"; +by (EVERY1 [etac image_eqI, rtac UNIV_I]); +bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI)); + +bind_thm ("rangeI", UNIV_I RS imageI); + +val [major,minor] = goal thy + "[| b : range(%x.f(x)); !!x. b=f(x) ==> P |] ==> P"; +by (rtac (major RS imageE) 1); +by (etac minor 1); +qed "rangeE"; + +AddIs [rangeI]; +AddSEs [rangeE]; + (*** Set reasoning tools ***) diff -r 8a680e310f04 -r 3fac3e8d5d3e src/HOL/Set.thy --- a/src/HOL/Set.thy Fri Apr 04 16:27:39 1997 +0200 +++ b/src/HOL/Set.thy Fri Apr 04 16:33:28 1997 +0200 @@ -32,8 +32,6 @@ Pow :: 'a set => 'a set set (*powerset*) range :: ('a => 'b) => 'b set (*of function*) Ball, Bex :: ['a set, 'a => bool] => bool (*bounded quantifiers*) - inj, surj :: ('a => 'b) => bool (*inj/surjective*) - inj_onto :: ['a => 'b, 'a set] => bool "``" :: ['a => 'b, 'a set] => ('b set) (infixr 90) (*membership*) "op :" :: ['a, 'a set] => bool ("(_/ : _)" [50, 51] 50) @@ -149,9 +147,6 @@ empty_def "{} == {x. False}" insert_def "insert a B == {x.x=a} Un B" image_def "f``A == {y. ? x:A. y=f(x)}" - inj_def "inj f == ! x y. f(x)=f(y) --> x=y" - inj_onto_def "inj_onto f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" - surj_def "surj f == ! y. ? x. y=f(x)" end diff -r 8a680e310f04 -r 3fac3e8d5d3e src/HOL/equalities.ML --- a/src/HOL/equalities.ML Fri Apr 04 16:27:39 1997 +0200 +++ b/src/HOL/equalities.ML Fri Apr 04 16:33:28 1997 +0200 @@ -412,7 +412,7 @@ by (Blast_tac 1); qed "Int_Union"; -(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: +(* Devlin, Setdamentals of Contemporary Set Theory, page 12, exercise 5: Union of a family of unions **) goal Set.thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; by (Blast_tac 1);