src/HOL/Real/HahnBanach/Linearform.thy
changeset 29197 6d4cb27ed19c
parent 29189 ee8572f3bb57
child 29198 418ed6411847
equal deleted inserted replaced
29189:ee8572f3bb57 29197:6d4cb27ed19c
     1 (*  Title:      HOL/Real/HahnBanach/Linearform.thy
       
     2     ID:         $Id$
       
     3     Author:     Gertrud Bauer, TU Munich
       
     4 *)
       
     5 
       
     6 header {* Linearforms *}
       
     7 
       
     8 theory Linearform
       
     9 imports VectorSpace
       
    10 begin
       
    11 
       
    12 text {*
       
    13   A \emph{linear form} is a function on a vector space into the reals
       
    14   that is additive and multiplicative.
       
    15 *}
       
    16 
       
    17 locale linearform = var V + var f +
       
    18   constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
       
    19   assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
       
    20     and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
       
    21 
       
    22 declare linearform.intro [intro?]
       
    23 
       
    24 lemma (in linearform) neg [iff]:
       
    25   assumes "vectorspace V"
       
    26   shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
       
    27 proof -
       
    28   interpret vectorspace [V] by fact
       
    29   assume x: "x \<in> V"
       
    30   then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
       
    31   also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
       
    32   also from x have "\<dots> = - (f x)" by simp
       
    33   finally show ?thesis .
       
    34 qed
       
    35 
       
    36 lemma (in linearform) diff [iff]:
       
    37   assumes "vectorspace V"
       
    38   shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
       
    39 proof -
       
    40   interpret vectorspace [V] by fact
       
    41   assume x: "x \<in> V" and y: "y \<in> V"
       
    42   then have "x - y = x + - y" by (rule diff_eq1)
       
    43   also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
       
    44   also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
       
    45   finally show ?thesis by simp
       
    46 qed
       
    47 
       
    48 text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
       
    49 
       
    50 lemma (in linearform) zero [iff]:
       
    51   assumes "vectorspace V"
       
    52   shows "f 0 = 0"
       
    53 proof -
       
    54   interpret vectorspace [V] by fact
       
    55   have "f 0 = f (0 - 0)" by simp
       
    56   also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
       
    57   also have "\<dots> = 0" by simp
       
    58   finally show ?thesis .
       
    59 qed
       
    60 
       
    61 end