| author | wenzelm | 
| Sat, 18 Aug 2007 13:32:20 +0200 | |
| changeset 24319 | 944705cc79d2 | 
| parent 23378 | 1d138d6bb461 | 
| child 25762 | c03e9d04b3e4 | 
| permissions | -rw-r--r-- | 
| 7566 | 1 | (* Title: HOL/Real/HahnBanach/Linearform.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Gertrud Bauer, TU Munich | |
| 4 | *) | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 5 | |
| 9035 | 6 | header {* Linearforms *}
 | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 7 | |
| 16417 | 8 | theory Linearform imports VectorSpace begin | 
| 7917 | 9 | |
| 10687 | 10 | text {*
 | 
| 11 |   A \emph{linear form} is a function on a vector space into the reals
 | |
| 12 | that is additive and multiplicative. | |
| 13 | *} | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 14 | |
| 13515 | 15 | locale linearform = var V + var f + | 
| 16 | assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y" | |
| 17 | and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x" | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 18 | |
| 14254 
342634f38451
Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
 ballarin parents: 
13547diff
changeset | 19 | declare linearform.intro [intro?] | 
| 
342634f38451
Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
 ballarin parents: 
13547diff
changeset | 20 | |
| 13547 | 21 | lemma (in linearform) neg [iff]: | 
| 22 | includes vectorspace | |
| 23 | shows "x \<in> V \<Longrightarrow> f (- x) = - f x" | |
| 10687 | 24 | proof - | 
| 13515 | 25 | assume x: "x \<in> V" | 
| 26 | hence "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1) | |
| 27 | also from x have "... = (- 1) * (f x)" by (rule mult) | |
| 28 | also from x have "... = - (f x)" by simp | |
| 9035 | 29 | finally show ?thesis . | 
| 30 | qed | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 31 | |
| 13547 | 32 | lemma (in linearform) diff [iff]: | 
| 33 | includes vectorspace | |
| 34 | shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y" | |
| 9035 | 35 | proof - | 
| 13515 | 36 | assume x: "x \<in> V" and y: "y \<in> V" | 
| 37 | hence "x - y = x + - y" by (rule diff_eq1) | |
| 13547 | 38 | also have "f ... = f x + f (- y)" by (rule add) (simp_all add: x y) | 
| 23378 | 39 | also have "f (- y) = - f y" using `vectorspace V` y by (rule neg) | 
| 13515 | 40 | finally show ?thesis by simp | 
| 9035 | 41 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 42 | |
| 10687 | 43 | text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
 | 
| 7917 | 44 | |
| 13547 | 45 | lemma (in linearform) zero [iff]: | 
| 46 | includes vectorspace | |
| 47 | shows "f 0 = 0" | |
| 10687 | 48 | proof - | 
| 13515 | 49 | have "f 0 = f (0 - 0)" by simp | 
| 23378 | 50 | also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all | 
| 13515 | 51 | also have "\<dots> = 0" by simp | 
| 52 | finally show ?thesis . | |
| 10687 | 53 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 54 | |
| 10687 | 55 | end |