# HG changeset patch # User paulson # Date 838376224 -7200 # Node ID 00b4b69929458976e57205b6e2345674cfb64402 # Parent 67f49e8c4355134994878a424243e861a28eae43 Redefining "range" as a macro diff -r 67f49e8c4355 -r 00b4b6992945 src/HOL/Fun.ML --- a/src/HOL/Fun.ML Fri Jul 26 12:16:17 1996 +0200 +++ b/src/HOL/Fun.ML Fri Jul 26 12:17:04 1996 +0200 @@ -19,28 +19,14 @@ qed "apply_inverse"; -(*** Range of a function ***) +(*** Image of a set under a function ***) (*Frequently b does not have the syntactic form of f(x).*) -val [prem] = goalw Fun.thy [range_def] "b=f(x) ==> b : range(f)"; -by (EVERY1 [rtac CollectI, rtac exI, rtac prem]); -qed "range_eqI"; - -val rangeI = refl RS range_eqI; - -val [major,minor] = goalw Fun.thy [range_def] - "[| b : range(%x.f(x)); !!x. b=f(x) ==> P |] ==> P"; -by (rtac (major RS CollectD RS exE) 1); -by (etac minor 1); -qed "rangeE"; - -(*** Image of a set under a function ***) - 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"; -val imageI = refl RS image_eqI; +bind_thm ("imageI", refl RS image_eqI); (*The eta-expansion gives variable-name preservation.*) val major::prems = goalw Fun.thy [image_def] @@ -59,6 +45,19 @@ by (fast_tac (!claset addIs [imageI,UnCI] addSEs [imageE,UnE]) 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] @@ -181,4 +180,4 @@ val set_cs = !claset delrules [equalityI]; -fun cfast_tac prems = cut_facts_tac prems THEN' fast_tac (claset_of "Fun"); + diff -r 67f49e8c4355 -r 00b4b6992945 src/HOL/Set.thy --- a/src/HOL/Set.thy Fri Jul 26 12:16:17 1996 +0200 +++ b/src/HOL/Set.thy Fri Jul 26 12:17:04 1996 +0200 @@ -60,6 +60,7 @@ translations "UNIV" == "Compl {}" + "range(f)" == "f``UNIV" "x ~: y" == "~ (x : y)" "{x, xs}" == "insert x {xs}" "{x}" == "insert x {}" @@ -97,7 +98,6 @@ Pow_def "Pow(A) == {B. B <= A}" empty_def "{} == {x. False}" insert_def "insert a B == {x.x=a} Un B" - range_def "range(f) == {y. ? x. y=f(x)}" 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"