src/HOL/Probability/Discrete_Topology.thy
author hoelzl
Mon, 14 Jan 2013 17:29:04 +0100
changeset 50881 ae630bab13da
parent 50245 dea9363887a6
child 51000 c9adb50f74ad
permissions -rw-r--r--
renamed countable_basis_space to second_countable_topology
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50089
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     1
(*  Title:      HOL/Probability/Discrete_Topology.thy
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     2
    Author:     Fabian Immler, TU München
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     3
*)
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     4
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     5
theory Discrete_Topology
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     6
imports Multivariate_Analysis
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     7
begin
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     8
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
     9
text {* Copy of discrete types with discrete topology. This space is polish. *}
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    10
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    11
typedef 'a discrete = "UNIV::'a set"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    12
morphisms of_discrete discrete
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    13
..
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    14
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    15
instantiation discrete :: (type) topological_space
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    16
begin
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    17
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    18
definition open_discrete::"'a discrete set \<Rightarrow> bool"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    19
  where "open_discrete s = True"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    20
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    21
instance proof qed (auto simp: open_discrete_def)
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    22
end
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    23
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    24
instantiation discrete :: (type) metric_space
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    25
begin
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    26
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    27
definition dist_discrete::"'a discrete \<Rightarrow> 'a discrete \<Rightarrow> real"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    28
  where "dist_discrete n m = (if n = m then 0 else 1)"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    29
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    30
instance proof qed (auto simp: open_discrete_def dist_discrete_def intro: exI[where x=1])
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    31
end
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    32
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    33
instance discrete :: (type) complete_space
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    34
proof
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    35
  fix X::"nat\<Rightarrow>'a discrete" assume "Cauchy X"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    36
  hence "\<exists>n. \<forall>m\<ge>n. X n = X m"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    37
    by (force simp: dist_discrete_def Cauchy_def split: split_if_asm dest:spec[where x=1])
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    38
  then guess n ..
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    39
  thus "convergent X"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    40
    by (intro convergentI[where L="X n"] tendstoI eventually_sequentiallyI[of n])
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    41
       (simp add: dist_discrete_def)
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    42
qed
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    43
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    44
instance discrete :: (countable) countable
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    45
proof
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    46
  have "inj (\<lambda>c::'a discrete. to_nat (of_discrete c))"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    47
    by (simp add: inj_on_def of_discrete_inject)
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    48
  thus "\<exists>f::'a discrete \<Rightarrow> nat. inj f" by blast
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    49
qed
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    50
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50245
diff changeset
    51
instance discrete :: (countable) second_countable_topology
50089
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    52
proof
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50089
diff changeset
    53
  let ?B = "(range (\<lambda>n::nat. {from_nat n::'a discrete}))"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50089
diff changeset
    54
  have "topological_basis ?B"
50089
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    55
  proof (intro topological_basisI)
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    56
    fix x::"'a discrete" and O' assume "open O'" "x \<in> O'"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    57
    thus "\<exists>B'\<in>range (\<lambda>n. {from_nat n}). x \<in> B' \<and> B' \<subseteq> O'"
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    58
      by (auto intro: exI[where x="to_nat x"])
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    59
  qed (simp add: open_discrete_def)
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50089
diff changeset
    60
  moreover have "countable ?B" by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50089
diff changeset
    61
  ultimately show "\<exists>B::'a discrete set set. countable B \<and> topological_basis B" by blast
50089
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    62
qed
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    63
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    64
instance discrete :: (countable) polish_space ..
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    65
1badf63e5d97 generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
diff changeset
    66
end