src/HOL/Probability/Fin_Map.thy
changeset 51473 1210309fddab
parent 51343 b61b32f62c78
child 51489 f738e6dbd844
equal deleted inserted replaced
51472:adb441e4b9e9 51473:1210309fddab
   220   qed
   220   qed
   221   then guess A unfolding choice_iff .. note A = this
   221   then guess A unfolding choice_iff .. note A = this
   222   hence open_sub: "\<And>i S. i\<in>domain x \<Longrightarrow> open (S i) \<Longrightarrow> x i\<in>(S i) \<Longrightarrow> (\<exists>a\<in>A i. a\<subseteq>(S i))" by auto
   222   hence open_sub: "\<And>i S. i\<in>domain x \<Longrightarrow> open (S i) \<Longrightarrow> x i\<in>(S i) \<Longrightarrow> (\<exists>a\<in>A i. a\<subseteq>(S i))" by auto
   223   have A_notempty: "\<And>i. i \<in> domain x \<Longrightarrow> A i \<noteq> {}" using open_sub[of _ "\<lambda>_. UNIV"] by auto
   223   have A_notempty: "\<And>i. i \<in> domain x \<Longrightarrow> A i \<noteq> {}" using open_sub[of _ "\<lambda>_. UNIV"] by auto
   224   let ?A = "(\<lambda>f. Pi' (domain x) f) ` (Pi\<^isub>E (domain x) A)"
   224   let ?A = "(\<lambda>f. Pi' (domain x) f) ` (Pi\<^isub>E (domain x) A)"
   225   show "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
   225   show "\<exists>A::nat \<Rightarrow> ('a\<Rightarrow>\<^isub>F'b) set. (\<forall>i. x \<in> (A i) \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
   226   proof (rule exI[where x="?A"], safe)
   226   proof (rule first_countableI[where A="?A"], safe)
   227     show "countable ?A" using A by (simp add: countable_PiE)
   227     show "countable ?A" using A by (simp add: countable_PiE)
   228   next
   228   next
   229     fix S::"('a \<Rightarrow>\<^isub>F 'b) set" assume "open S" "x \<in> S"
   229     fix S::"('a \<Rightarrow>\<^isub>F 'b) set" assume "open S" "x \<in> S"
   230     thus "\<exists>a\<in>?A. a \<subseteq> S" unfolding open_finmap_def
   230     thus "\<exists>a\<in>?A. a \<subseteq> S" unfolding open_finmap_def
   231     proof (induct rule: generate_topology.induct)
   231     proof (induct rule: generate_topology.induct)