diff -r 10080e31b294 -r c350f3ad6b0d src/HOL/RealVector.thy --- a/src/HOL/RealVector.thy Sat Jun 06 10:28:34 2009 -0700 +++ b/src/HOL/RealVector.thy Sun Jun 07 12:00:03 2009 -0700 @@ -421,10 +421,94 @@ class topo = fixes topo :: "'a set set" +text {* + The syntactic class uses @{term topo} instead of @{text "open"} + so that @{text "open"} and @{text closed} will have the same type. + Maybe we could use extra type constraints instead, like with + @{text dist} and @{text norm}? +*} + class topological_space = topo + assumes topo_UNIV: "UNIV \ topo" assumes topo_Int: "A \ topo \ B \ topo \ A \ B \ topo" assumes topo_Union: "T \ topo \ \T \ topo" +begin + +definition + "open" :: "'a set \ bool" where + "open S \ S \ topo" + +definition + closed :: "'a set \ bool" where + "closed S \ open (- S)" + +lemma open_UNIV [intro, simp]: "open UNIV" + unfolding open_def by (rule topo_UNIV) + +lemma open_Int [intro]: "open S \ open T \ open (S \ T)" + unfolding open_def by (rule topo_Int) + +lemma open_Union [intro]: "\S\K. open S \ open (\ K)" + unfolding open_def subset_eq [symmetric] by (rule topo_Union) + +lemma open_empty [intro, simp]: "open {}" + using open_Union [of "{}"] by simp + +lemma open_Un [intro]: "open S \ open T \ open (S \ T)" + using open_Union [of "{S, T}"] by simp + +lemma open_UN [intro]: "\x\A. open (B x) \ open (\x\A. B x)" + unfolding UN_eq by (rule open_Union) auto + +lemma open_INT [intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" + by (induct set: finite) auto + +lemma open_Inter [intro]: "finite S \ \T\S. open T \ open (\S)" + unfolding Inter_def by (rule open_INT) + +lemma closed_empty [intro, simp]: "closed {}" + unfolding closed_def by simp + +lemma closed_Un [intro]: "closed S \ closed T \ closed (S \ T)" + unfolding closed_def by auto + +lemma closed_Inter [intro]: "\S\K. closed S \ closed (\ K)" + unfolding closed_def Inter_def by auto + +lemma closed_UNIV [intro, simp]: "closed UNIV" + unfolding closed_def by simp + +lemma closed_Int [intro]: "closed S \ closed T \ closed (S \ T)" + unfolding closed_def by auto + +lemma closed_INT [intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" + unfolding closed_def by auto + +lemma closed_UN [intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" + by (induct set: finite) auto + +lemma closed_Union [intro]: "finite S \ \T\S. closed T \ closed (\S)" + unfolding Union_def by (rule closed_UN) + +lemma open_closed: "open S \ closed (- S)" + unfolding closed_def by simp + +lemma closed_open: "closed S \ open (- S)" + unfolding closed_def by simp + +lemma open_Diff [intro]: "open S \ closed T \ open (S - T)" + unfolding closed_open Diff_eq by (rule open_Int) + +lemma closed_Diff [intro]: "closed S \ open T \ closed (S - T)" + unfolding open_closed Diff_eq by (rule closed_Int) + +lemma open_Compl [intro]: "closed S \ open (- S)" + unfolding closed_open . + +lemma closed_Compl [intro]: "open S \ closed (- S)" + unfolding open_closed . + +end subsection {* Metric spaces *}