diff -r f4c079225845 -r c12b25b7f015 src/HOL/Library/Product_Vector.thy --- a/src/HOL/Library/Product_Vector.thy Wed Jun 03 08:46:13 2009 -0700 +++ b/src/HOL/Library/Product_Vector.thy Wed Jun 03 09:58:11 2009 -0700 @@ -45,28 +45,29 @@ "*" :: (topological_space, topological_space) topological_space begin -definition open_prod_def: - "open S = (\x\S. \A B. open A \ open B \ x \ A \ B \ A \ B \ S)" +definition topo_prod_def: + "topo = {S. \x\S. \A\topo. \B\topo. x \ A \ B \ A \ B \ S}" instance proof - show "open (UNIV :: ('a \ 'b) set)" - unfolding open_prod_def by (fast intro: open_UNIV) + show "(UNIV :: ('a \ 'b) set) \ topo" + unfolding topo_prod_def by (auto intro: topo_UNIV) next fix S T :: "('a \ 'b) set" - assume "open S" "open T" thus "open (S \ T)" - unfolding open_prod_def + assume "S \ topo" "T \ topo" thus "S \ T \ topo" + unfolding topo_prod_def apply clarify apply (drule (1) bspec)+ apply (clarify, rename_tac Sa Ta Sb Tb) - apply (rule_tac x="Sa \ Ta" in exI) - apply (rule_tac x="Sb \ Tb" in exI) - apply (simp add: open_Int) + apply (rule_tac x="Sa \ Ta" in rev_bexI) + apply (simp add: topo_Int) + apply (rule_tac x="Sb \ Tb" in rev_bexI) + apply (simp add: topo_Int) apply fast done next fix T :: "('a \ 'b) set set" - assume "\A\T. open A" thus "open (\T)" - unfolding open_prod_def by fast + assume "T \ topo" thus "\T \ topo" + unfolding topo_prod_def Bex_def by fast qed end @@ -103,10 +104,9 @@ (* FIXME: long proof! *) (* Maybe it would be easier to define topological spaces *) (* in terms of neighborhoods instead of open sets? *) - fix S :: "('a \ 'b) set" - show "open S = (\x\S. \e>0. \y. dist y x < e \ y \ S)" - unfolding open_prod_def open_dist - apply safe + show "topo = {S::('a \ 'b) set. \x\S. \e>0. \y. dist y x < e \ y \ S}" + unfolding topo_prod_def topo_dist + apply (safe, rename_tac S a b) apply (drule (1) bspec) apply clarify apply (drule (1) bspec)+ @@ -121,19 +121,18 @@ apply (drule spec, erule mp) apply (erule le_less_trans [OF real_sqrt_sum_squares_ge2]) + apply (rename_tac S a b) apply (drule (1) bspec) apply clarify apply (subgoal_tac "\r>0. \s>0. e = sqrt (r\ + s\)") apply clarify - apply (rule_tac x="{y. dist y a < r}" in exI) - apply (rule_tac x="{y. dist y b < s}" in exI) - apply (rule conjI) + apply (rule_tac x="{y. dist y a < r}" in rev_bexI) apply clarify apply (rule_tac x="r - dist x a" in exI, rule conjI, simp) apply clarify apply (rule le_less_trans [OF dist_triangle]) apply (erule less_le_trans [OF add_strict_right_mono], simp) - apply (rule conjI) + apply (rule_tac x="{y. dist y b < s}" in rev_bexI) apply clarify apply (rule_tac x="s - dist x b" in exI, rule conjI, simp) apply clarify