# HG changeset patch # User immler # Date 1352991001 -3600 # Node ID 1badf63e5d977cc45ca96c1883090a7d90ffc958 # Parent 32d1795cc77a70e8683a3056f8d6d3a69f883a20 generalized to copy of countable types instead of instantiation of nat for discrete topology diff -r 32d1795cc77a -r 1badf63e5d97 src/HOL/Probability/Discrete_Topology.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Discrete_Topology.thy Thu Nov 15 15:50:01 2012 +0100 @@ -0,0 +1,64 @@ +(* Title: HOL/Probability/Discrete_Topology.thy + Author: Fabian Immler, TU München +*) + +theory Discrete_Topology +imports Multivariate_Analysis +begin + +text {* Copy of discrete types with discrete topology. This space is polish. *} + +typedef 'a discrete = "UNIV::'a set" +morphisms of_discrete discrete +.. + +instantiation discrete :: (type) topological_space +begin + +definition open_discrete::"'a discrete set \ bool" + where "open_discrete s = True" + +instance proof qed (auto simp: open_discrete_def) +end + +instantiation discrete :: (type) metric_space +begin + +definition dist_discrete::"'a discrete \ 'a discrete \ real" + where "dist_discrete n m = (if n = m then 0 else 1)" + +instance proof qed (auto simp: open_discrete_def dist_discrete_def intro: exI[where x=1]) +end + +instance discrete :: (type) complete_space +proof + fix X::"nat\'a discrete" assume "Cauchy X" + hence "\n. \m\n. X n = X m" + by (force simp: dist_discrete_def Cauchy_def split: split_if_asm dest:spec[where x=1]) + then guess n .. + thus "convergent X" + by (intro convergentI[where L="X n"] tendstoI eventually_sequentiallyI[of n]) + (simp add: dist_discrete_def) +qed + +instance discrete :: (countable) countable +proof + have "inj (\c::'a discrete. to_nat (of_discrete c))" + by (simp add: inj_on_def of_discrete_inject) + thus "\f::'a discrete \ nat. inj f" by blast +qed + +instance discrete :: (countable) enumerable_basis +proof + have "topological_basis (range (\n::nat. {from_nat n::'a discrete}))" + proof (intro topological_basisI) + fix x::"'a discrete" and O' assume "open O'" "x \ O'" + thus "\B'\range (\n. {from_nat n}). x \ B' \ B' \ O'" + by (auto intro: exI[where x="to_nat x"]) + qed (simp add: open_discrete_def) + thus "\f::nat\'a discrete set. topological_basis (range f)" by blast +qed + +instance discrete :: (countable) polish_space .. + +end diff -r 32d1795cc77a -r 1badf63e5d97 src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Thu Nov 15 11:16:58 2012 +0100 +++ b/src/HOL/Probability/Probability.thy Thu Nov 15 15:50:01 2012 +0100 @@ -1,5 +1,6 @@ theory Probability imports + Discrete_Topology Complete_Measure Probability_Measure Infinite_Product_Measure diff -r 32d1795cc77a -r 1badf63e5d97 src/HOL/Probability/Regularity.thy --- a/src/HOL/Probability/Regularity.thy Thu Nov 15 11:16:58 2012 +0100 +++ b/src/HOL/Probability/Regularity.thy Thu Nov 15 15:50:01 2012 +0100 @@ -2,53 +2,12 @@ Author: Fabian Immler, TU München *) +header {* Regularity of Measures *} + theory Regularity imports Measure_Space Borel_Space begin -instantiation nat::topological_space -begin - -definition open_nat::"nat set \ bool" - where "open_nat s = True" - -instance proof qed (auto simp: open_nat_def) -end - -instantiation nat::metric_space -begin - -definition dist_nat::"nat \ nat \ real" - where "dist_nat n m = (if n = m then 0 else 1)" - -instance proof qed (auto simp: open_nat_def dist_nat_def intro: exI[where x=1]) -end - -instance nat::complete_space -proof - fix X::"nat\nat" assume "Cauchy X" - hence "\n. \m\n. X m = X n" - by (force simp: dist_nat_def Cauchy_def split: split_if_asm dest:spec[where x=1]) - then guess n .. - thus "convergent X" - apply (intro convergentI[where L="X n"] tendstoI) - unfolding eventually_sequentially dist_nat_def - apply (intro exI[where x=n]) - apply (intro allI) - apply (drule_tac x=na in spec) - apply simp - done -qed - -instance nat::enumerable_basis -proof - have "topological_basis (range (\n::nat. {n}))" - by (intro topological_basisI) (auto simp: open_nat_def) - thus "\f::nat\nat set. topological_basis (range f)" by blast -qed - -subsection {* Regularity of Measures *} - lemma ereal_approx_SUP: fixes x::ereal assumes A_notempty: "A \ {}"