src/HOL/Induct/Sigma_Algebra.thy
author wenzelm
Fri, 08 Jan 2021 22:30:32 +0100
changeset 73110 c87ca43ebd3b
parent 63233 e53830948c4f
permissions -rw-r--r--
avoid rescaled fonts, e.g. dockable buttons on Windows L&F after opening a new view;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Induct/Sigma_Algebra.thy
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     3
*)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     4
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     5
section \<open>Sigma algebras\<close>
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 10875
diff changeset
     6
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
     7
theory Sigma_Algebra
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
     8
imports Main
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
     9
begin
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    10
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    11
text \<open>
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    12
  This is just a tiny example demonstrating the use of inductive
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63067
diff changeset
    13
  definitions in classical mathematics.  We define the least \<open>\<sigma>\<close>-algebra over a given set of sets.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    14
\<close>
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    15
63233
wenzelm
parents: 63167
diff changeset
    16
inductive_set \<sigma>_algebra :: "'a set set \<Rightarrow> 'a set set" for A :: "'a set set"
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
    17
where
63233
wenzelm
parents: 63167
diff changeset
    18
  basic: "a \<in> \<sigma>_algebra A" if "a \<in> A" for a
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
    19
| UNIV: "UNIV \<in> \<sigma>_algebra A"
63233
wenzelm
parents: 63167
diff changeset
    20
| complement: "- a \<in> \<sigma>_algebra A" if "a \<in> \<sigma>_algebra A" for a
wenzelm
parents: 63167
diff changeset
    21
| Union: "(\<Union>i. a i) \<in> \<sigma>_algebra A" if "\<And>i::nat. a i \<in> \<sigma>_algebra A" for a
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    22
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    23
text \<open>
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    24
  The following basic facts are consequences of the closure properties
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63067
diff changeset
    25
  of any \<open>\<sigma>\<close>-algebra, merely using the introduction rules, but
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    26
  no induction nor cases.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    27
\<close>
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    28
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    29
theorem sigma_algebra_empty: "{} \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    30
proof -
14717
7d8d4c9b36fd tuned notation;
wenzelm
parents: 11046
diff changeset
    31
  have "UNIV \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.UNIV)
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
    32
  then have "-UNIV \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.complement)
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    33
  also have "-UNIV = {}" by simp
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    34
  finally show ?thesis .
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    35
qed
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    36
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    37
theorem sigma_algebra_Inter:
60532
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    38
  "(\<And>i::nat. a i \<in> \<sigma>_algebra A) \<Longrightarrow> (\<Inter>i. a i) \<in> \<sigma>_algebra A"
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    39
proof -
60532
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    40
  assume "\<And>i::nat. a i \<in> \<sigma>_algebra A"
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    41
  then have "\<And>i::nat. -(a i) \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.complement)
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
    42
  then have "(\<Union>i. -(a i)) \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.Union)
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 36862
diff changeset
    43
  then have "-(\<Union>i. -(a i)) \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.complement)
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    44
  also have "-(\<Union>i. -(a i)) = (\<Inter>i. a i)" by simp
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    45
  finally show ?thesis .
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    46
qed
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    47
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    48
end