| author | haftmann | 
| Wed, 23 Sep 2009 11:33:52 +0200 | |
| changeset 32647 | e54f47f9e28b | 
| parent 23746 | a455e69c31cc | 
| child 36862 | 952b2b102a0a | 
| permissions | -rw-r--r-- | 
| 10875 | 1 | (* Title: HOL/Induct/Sigma_Algebra.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 11046 
b5f5942781a0
Induct: converted some theories to new-style format;
 wenzelm parents: 
10875diff
changeset | 6 | header {* Sigma algebras *}
 | 
| 
b5f5942781a0
Induct: converted some theories to new-style format;
 wenzelm parents: 
10875diff
changeset | 7 | |
| 16417 | 8 | theory Sigma_Algebra imports Main begin | 
| 10875 | 9 | |
| 10 | text {*
 | |
| 11 | This is just a tiny example demonstrating the use of inductive | |
| 12 |   definitions in classical mathematics.  We define the least @{text
 | |
| 13 | \<sigma>}-algebra over a given set of sets. | |
| 14 | *} | |
| 15 | ||
| 23746 | 16 | inductive_set | 
| 14717 | 17 | \<sigma>_algebra :: "'a set set => 'a set set" | 
| 23746 | 18 | for A :: "'a set set" | 
| 19 | where | |
| 10875 | 20 | basic: "a \<in> A ==> a \<in> \<sigma>_algebra A" | 
| 23746 | 21 | | UNIV: "UNIV \<in> \<sigma>_algebra A" | 
| 22 | | complement: "a \<in> \<sigma>_algebra A ==> -a \<in> \<sigma>_algebra A" | |
| 23 | | Union: "(!!i::nat. a i \<in> \<sigma>_algebra A) ==> (\<Union>i. a i) \<in> \<sigma>_algebra A" | |
| 10875 | 24 | |
| 25 | text {*
 | |
| 26 | The following basic facts are consequences of the closure properties | |
| 27 |   of any @{text \<sigma>}-algebra, merely using the introduction rules, but
 | |
| 28 | no induction nor cases. | |
| 29 | *} | |
| 30 | ||
| 31 | theorem sigma_algebra_empty: "{} \<in> \<sigma>_algebra A"
 | |
| 32 | proof - | |
| 14717 | 33 | have "UNIV \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.UNIV) | 
| 34 | hence "-UNIV \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.complement) | |
| 10875 | 35 |   also have "-UNIV = {}" by simp
 | 
| 36 | finally show ?thesis . | |
| 37 | qed | |
| 38 | ||
| 39 | theorem sigma_algebra_Inter: | |
| 40 | "(!!i::nat. a i \<in> \<sigma>_algebra A) ==> (\<Inter>i. a i) \<in> \<sigma>_algebra A" | |
| 41 | proof - | |
| 42 | assume "!!i::nat. a i \<in> \<sigma>_algebra A" | |
| 14717 | 43 | hence "!!i::nat. -(a i) \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.complement) | 
| 44 | hence "(\<Union>i. -(a i)) \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.Union) | |
| 45 | hence "-(\<Union>i. -(a i)) \<in> \<sigma>_algebra A" by (rule \<sigma>_algebra.complement) | |
| 10875 | 46 | also have "-(\<Union>i. -(a i)) = (\<Inter>i. a i)" by simp | 
| 47 | finally show ?thesis . | |
| 48 | qed | |
| 49 | ||
| 50 | end |