src/HOL/Real/HahnBanach/Linearform.thy
author wenzelm
Tue, 16 May 2006 13:01:22 +0200
changeset 19640 40ec89317425
parent 16417 9bc16273c2d4
child 23378 1d138d6bb461
permissions -rw-r--r--
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/Linearform.thy
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    ID:         $Id$
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     4
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     5
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
     6
header {* Linearforms *}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14254
diff changeset
     8
theory Linearform imports VectorSpace begin
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
     9
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    10
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    11
  A \emph{linear form} is a function on a vector space into the reals
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    12
  that is additive and multiplicative.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    13
*}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    14
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    15
locale linearform = var V + var f +
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    16
  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    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: 13547
diff changeset
    19
declare linearform.intro [intro?]
342634f38451 Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents: 13547
diff changeset
    20
13547
wenzelm
parents: 13515
diff changeset
    21
lemma (in linearform) neg [iff]:
wenzelm
parents: 13515
diff changeset
    22
  includes vectorspace
wenzelm
parents: 13515
diff changeset
    23
  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    24
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    25
  assume x: "x \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    26
  hence "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    27
  also from x have "... = (- 1) * (f x)" by (rule mult)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    28
  also from x have "... = - (f x)" by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    29
  finally show ?thesis .
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    30
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    31
13547
wenzelm
parents: 13515
diff changeset
    32
lemma (in linearform) diff [iff]:
wenzelm
parents: 13515
diff changeset
    33
  includes vectorspace
wenzelm
parents: 13515
diff changeset
    34
  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    35
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    36
  assume x: "x \<in> V" and y: "y \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    37
  hence "x - y = x + - y" by (rule diff_eq1)
13547
wenzelm
parents: 13515
diff changeset
    38
  also have "f ... = f x + f (- y)" by (rule add) (simp_all add: x y)
wenzelm
parents: 13515
diff changeset
    39
  also from _ y have "f (- y) = - f y" by (rule neg)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    40
  finally show ?thesis by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    41
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    42
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    43
text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    44
13547
wenzelm
parents: 13515
diff changeset
    45
lemma (in linearform) zero [iff]:
wenzelm
parents: 13515
diff changeset
    46
  includes vectorspace
wenzelm
parents: 13515
diff changeset
    47
  shows "f 0 = 0"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    48
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    49
  have "f 0 = f (0 - 0)" by simp
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    50
  also have "\<dots> = f 0 - f 0" by (rule diff) simp_all
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    51
  also have "\<dots> = 0" by simp
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    52
  finally show ?thesis .
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    53
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    54
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9408
diff changeset
    55
end