src/HOL/Induct/Sigma_Algebra.thy
author wenzelm
Fri, 05 Oct 2001 21:52:39 +0200
changeset 11701 3d51fbf81c17
parent 11046 b5f5942781a0
child 14717 7d8d4c9b36fd
permissions -rw-r--r--
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat, "num" syntax (still with "#"), Numeral0, Numeral1;
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
    ID:         $Id$
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     5
*)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     6
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 10875
diff changeset
     7
header {* Sigma algebras *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 10875
diff changeset
     8
10875
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
     9
theory Sigma_Algebra = Main:
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    10
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    11
text {*
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    12
  This is just a tiny example demonstrating the use of inductive
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    13
  definitions in classical mathematics.  We define the least @{text
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    14
  \<sigma>}-algebra over a given set of sets.
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    15
*}
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    16
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    17
consts
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    18
  sigma_algebra :: "'a set set => 'a set set"    ("\<sigma>'_algebra")
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    19
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    20
inductive "\<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    21
  intros
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    22
    basic: "a \<in> A ==> a \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    23
    UNIV: "UNIV \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    24
    complement: "a \<in> \<sigma>_algebra A ==> -a \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    25
    Union: "(!!i::nat. a i \<in> \<sigma>_algebra A) ==> (\<Union>i. a i) \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    26
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    27
text {*
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    28
  The following basic facts are consequences of the closure properties
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    29
  of any @{text \<sigma>}-algebra, merely using the introduction rules, but
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    30
  no induction nor cases.
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    31
*}
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    32
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    33
theorem sigma_algebra_empty: "{} \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    34
proof -
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    35
  have "UNIV \<in> \<sigma>_algebra A" by (rule sigma_algebra.UNIV)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    36
  hence "-UNIV \<in> \<sigma>_algebra A" by (rule sigma_algebra.complement)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    37
  also have "-UNIV = {}" by simp
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    38
  finally show ?thesis .
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    39
qed
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    40
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    41
theorem sigma_algebra_Inter:
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    42
  "(!!i::nat. a i \<in> \<sigma>_algebra A) ==> (\<Inter>i. a i) \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    43
proof -
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    44
  assume "!!i::nat. a i \<in> \<sigma>_algebra A"
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    45
  hence "!!i::nat. -(a i) \<in> \<sigma>_algebra A" by (rule sigma_algebra.complement)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    46
  hence "(\<Union>i. -(a i)) \<in> \<sigma>_algebra A" by (rule sigma_algebra.Union)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    47
  hence "-(\<Union>i. -(a i)) \<in> \<sigma>_algebra A" by (rule sigma_algebra.complement)
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    48
  also have "-(\<Union>i. -(a i)) = (\<Inter>i. a i)" by simp
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    49
  finally show ?thesis .
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    50
qed
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    51
1715cb147294 added Induct/Sigma_Algebra.thy;
wenzelm
parents:
diff changeset
    52
end