src/HOL/Analysis/Finite_Cartesian_Product.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81472 e43bff789ac0
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63040
diff changeset
     1
(*  Title:      HOL/Analysis/Finite_Cartesian_Product.thy
35253
68dd8b51c6b8 tuned headers;
wenzelm
parents: 35113
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     3
*)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     4
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
     5
section \<open>Definition of Finite Cartesian Product Type\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     6
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     7
theory Finite_Cartesian_Product
41413
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 39302
diff changeset
     8
imports
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
     9
  Euclidean_Space
41413
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 39302
diff changeset
    10
  L2_Norm
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66281
diff changeset
    11
  "HOL-Library.Numeral_Type"
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66281
diff changeset
    12
  "HOL-Library.Countable_Set"
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
    13
  "HOL-Library.FuncSet"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    14
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    15
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
    16
subsection\<^marker>\<open>tag unimportant\<close> \<open>Finite Cartesian products, with indexing and lambdas\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    17
67731
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
    18
typedef ('a, 'b) vec = "UNIV :: ('b::finite \<Rightarrow> 'a) set"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    19
  morphisms vec_nth vec_lambda ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    20
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    21
declare vec_lambda_inject [simplified, simp]
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    22
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81100
diff changeset
    23
open_bundle vec_syntax
81100
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 80914
diff changeset
    24
begin
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 80914
diff changeset
    25
notation vec_nth (infixl \<open>$\<close> 90) and vec_lambda (binder \<open>\<chi>\<close> 10)
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67673
diff changeset
    26
end
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67673
diff changeset
    27
67731
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
    28
text \<open>
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
    29
  Concrete syntax for \<open>('a, 'b) vec\<close>:
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
    30
    \<^item> \<open>'a^'b\<close> becomes \<open>('a, 'b::finite) vec\<close>
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
    31
    \<^item> \<open>'a^'b::_\<close> becomes \<open>('a, 'b) vec\<close> without extra sort-constraint
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
    32
\<close>
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80768
diff changeset
    33
syntax "_vec_type" :: "type \<Rightarrow> type \<Rightarrow> type" (infixl \<open>^\<close> 15)
80768
c7723cc15de8 more markup for syntax consts;
wenzelm
parents: 73976
diff changeset
    34
syntax_types "_vec_type" \<rightleftharpoons> vec
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
    35
parse_translation \<open>
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    36
  let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
    37
    fun vec t u = Syntax.const \<^type_syntax>\<open>vec\<close> $ t $ u;
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    38
    fun finite_vec_tr [t, u] =
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    39
      (case Term_Position.strip_positions u of
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    40
        v as Free (x, _) =>
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    41
          if Lexicon.is_tid x then
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
    42
            vec t (Syntax.const \<^syntax_const>\<open>_ofsort\<close> $ v $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
    43
              Syntax.const \<^class_syntax>\<open>finite\<close>)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    44
          else vec t u
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    45
      | _ => vec t u)
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    46
  in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
    47
    [(\<^syntax_const>\<open>_vec_type\<close>, K finite_vec_tr)]
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51002
diff changeset
    48
  end
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
    49
\<close>
34290
1edf0f223c6e added syntax translation to automatically add finite typeclass to index type of cartesian product type
hoelzl
parents: 34289
diff changeset
    50
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    51
lemma vec_eq_iff: "(x = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)"
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    52
  by (simp add: vec_nth_inject [symmetric] fun_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    53
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    54
lemma vec_lambda_beta [simp]: "vec_lambda g $ i = g i"
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    55
  by (simp add: vec_lambda_inverse)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    56
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    57
lemma vec_lambda_unique: "(\<forall>i. f$i = g i) \<longleftrightarrow> vec_lambda g = f"
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    58
  by (auto simp add: vec_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    59
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    60
lemma vec_lambda_eta [simp]: "(\<chi> i. (g$i)) = g"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
    61
  by (simp add: vec_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    62
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
    63
subsection \<open>Cardinality of vectors\<close>
66281
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    64
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    65
instance vec :: (finite, finite) finite
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    66
proof
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    67
  show "finite (UNIV :: ('a, 'b) vec set)"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    68
  proof (subst bij_betw_finite)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    69
    show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    70
      by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    71
    have "finite (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    72
      by (intro finite_PiE) auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    73
    also have "(PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set)) = Pi UNIV (\<lambda>_. UNIV)"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    74
      by auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    75
    finally show "finite \<dots>" .
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    76
  qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    77
qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    78
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    79
lemma countable_PiE:
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    80
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (Pi\<^sub>E I F)"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    81
  by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    82
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    83
instance vec :: (countable, finite) countable
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    84
proof
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    85
  have "countable (UNIV :: ('a, 'b) vec set)"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    86
  proof (rule countableI_bij2)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    87
    show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    88
      by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    89
    have "countable (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    90
      by (intro countable_PiE) auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    91
    also have "(PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set)) = Pi UNIV (\<lambda>_. UNIV)"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    92
      by auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    93
    finally show "countable \<dots>" .
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    94
  qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    95
  thus "\<exists>t::('a, 'b) vec \<Rightarrow> nat. inj t"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    96
    by (auto elim!: countableE)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    97
qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    98
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
    99
lemma infinite_UNIV_vec:
66281
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   100
  assumes "infinite (UNIV :: 'a set)"
67731
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
   101
  shows   "infinite (UNIV :: ('a^'b) set)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   102
proof (subst bij_betw_finite)
66281
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   103
  show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   104
    by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   105
  have "infinite (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" (is "infinite ?A")
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   106
  proof
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   107
    assume "finite ?A"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   108
    hence "finite ((\<lambda>f. f undefined) ` ?A)"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   109
      by (rule finite_imageI)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   110
    also have "(\<lambda>f. f undefined) ` ?A = UNIV"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   111
      by auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   112
    finally show False 
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   113
      using \<open>infinite (UNIV :: 'a set)\<close> by contradiction
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   114
  qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   115
  also have "?A = Pi UNIV (\<lambda>_. UNIV)" 
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   116
    by auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   117
  finally show "infinite (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" .
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   118
qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   119
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   120
proposition CARD_vec [simp]:
67731
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
   121
  "CARD('a^'b) = CARD('a) ^ CARD('b)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   122
proof (cases "finite (UNIV :: 'a set)")
66281
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   123
  case True
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   124
  show ?thesis
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   125
  proof (subst bij_betw_same_card)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   126
    show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   127
      by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff)
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   128
    have "CARD('a) ^ CARD('b) = card (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   129
      (is "_ = card ?A")
71044
nipkow
parents: 70136
diff changeset
   130
      by (subst card_PiE) (auto)
66281
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   131
    also have "?A = Pi UNIV (\<lambda>_. UNIV)" 
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   132
      by auto
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   133
    finally show "card \<dots> = CARD('a) ^ CARD('b)" ..
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   134
  qed
6ad54b84ca5d facts about cardinality of vector type
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
   135
qed (simp_all add: infinite_UNIV_vec)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   136
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   137
lemma countable_vector:
67683
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   138
  fixes B:: "'n::finite \<Rightarrow> 'a set"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   139
  assumes "\<And>i. countable (B i)"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   140
  shows "countable {V. \<forall>i::'n::finite. V $ i \<in> B i}"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   141
proof -
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   142
  have "f \<in> ($) ` {V. \<forall>i. V $ i \<in> B i}" if "f \<in> Pi\<^sub>E UNIV B" for f
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   143
  proof -
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   144
    have "\<exists>W. (\<forall>i. W $ i \<in> B i) \<and> ($) W = f"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   145
      by (metis that PiE_iff UNIV_I vec_lambda_inverse)
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   146
    then show "f \<in> ($) ` {v. \<forall>i. v $ i \<in> B i}"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   147
      by blast
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   148
  qed
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   149
  then have "Pi\<^sub>E UNIV B = vec_nth ` {V. \<forall>i::'n. V $ i \<in> B i}"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   150
    by blast
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   151
  then have "countable (vec_nth ` {V. \<forall>i. V $ i \<in> B i})"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   152
    by (metis finite_class.finite_UNIV countable_PiE assms)
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   153
  then have "countable (vec_lambda ` vec_nth ` {V. \<forall>i. V $ i \<in> B i})"
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   154
    by auto
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   155
  then show ?thesis
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   156
    by (simp add: image_comp o_def vec_nth_inverse)
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   157
qed
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   158
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   159
subsection\<^marker>\<open>tag unimportant\<close> \<open>Group operations and class instances\<close>
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   160
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   161
instantiation vec :: (zero, finite) zero
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   162
begin
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   163
  definition "0 \<equiv> (\<chi> i. 0)"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   164
  instance ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   165
end
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   166
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   167
instantiation vec :: (plus, finite) plus
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   168
begin
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
   169
  definition "(+) \<equiv> (\<lambda> x y. (\<chi> i. x$i + y$i))"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   170
  instance ..
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   171
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   172
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   173
instantiation vec :: (minus, finite) minus
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   174
begin
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
   175
  definition "(-) \<equiv> (\<lambda> x y. (\<chi> i. x$i - y$i))"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   176
  instance ..
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   177
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   178
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   179
instantiation vec :: (uminus, finite) uminus
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   180
begin
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   181
  definition "uminus \<equiv> (\<lambda> x. (\<chi> i. - (x$i)))"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   182
  instance ..
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   183
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   184
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   185
lemma zero_index [simp]: "0 $ i = 0"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   186
  unfolding zero_vec_def by simp
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   187
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   188
lemma vector_add_component [simp]: "(x + y)$i = x$i + y$i"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   189
  unfolding plus_vec_def by simp
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   190
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   191
lemma vector_minus_component [simp]: "(x - y)$i = x$i - y$i"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   192
  unfolding minus_vec_def by simp
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   193
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   194
lemma vector_uminus_component [simp]: "(- x)$i = - (x$i)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   195
  unfolding uminus_vec_def by simp
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   196
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   197
instance vec :: (semigroup_add, finite) semigroup_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   198
  by standard (simp add: vec_eq_iff add.assoc)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   199
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   200
instance vec :: (ab_semigroup_add, finite) ab_semigroup_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   201
  by standard (simp add: vec_eq_iff add.commute)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   202
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   203
instance vec :: (monoid_add, finite) monoid_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   204
  by standard (simp_all add: vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   205
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   206
instance vec :: (comm_monoid_add, finite) comm_monoid_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   207
  by standard (simp add: vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   208
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   209
instance vec :: (cancel_semigroup_add, finite) cancel_semigroup_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   210
  by standard (simp_all add: vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   211
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   212
instance vec :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   213
  by standard (simp_all add: vec_eq_iff diff_diff_eq)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   214
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   215
instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   216
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   217
instance vec :: (group_add, finite) group_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   218
  by standard (simp_all add: vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   219
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   220
instance vec :: (ab_group_add, finite) ab_group_add
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   221
  by standard (simp_all add: vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   222
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   223
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   224
subsection\<^marker>\<open>tag unimportant\<close>\<open>Basic componentwise operations on vectors\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   225
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   226
instantiation vec :: (times, finite) times
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   227
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   228
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   229
definition "(*) \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   230
instance ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   231
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   232
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   233
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   234
instantiation vec :: (one, finite) one
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   235
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   236
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   237
definition "1 \<equiv> (\<chi> i. 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   238
instance ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   239
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   240
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   241
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   242
instantiation vec :: (ord, finite) ord
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   243
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   244
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   245
definition "x \<le> y \<longleftrightarrow> (\<forall>i. x$i \<le> y$i)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   246
definition "x < (y::'a^'b) \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   247
instance ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   248
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   249
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   250
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   251
text\<open>The ordering on one-dimensional vectors is linear.\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   252
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   253
instance vec:: (order, finite) order
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   254
  by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   255
      intro: order.trans order.antisym order.strict_implies_order)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   256
69680
96a43caa4902 revert to 56acd449da41
immler
parents: 69678
diff changeset
   257
instance vec :: (linorder, CARD_1) linorder
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   258
proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   259
  obtain a :: 'b where all: "\<And>P. (\<forall>i. P i) \<longleftrightarrow> P a"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   260
  proof -
69680
96a43caa4902 revert to 56acd449da41
immler
parents: 69678
diff changeset
   261
    have "CARD ('b) = 1" by (rule CARD_1)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   262
    then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   263
    then have "\<And>P. (\<forall>i\<in>UNIV. P i) \<longleftrightarrow> P b" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   264
    then show thesis by (auto intro: that)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   265
  qed
69680
96a43caa4902 revert to 56acd449da41
immler
parents: 69678
diff changeset
   266
  fix x y :: "'a^'b::CARD_1"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   267
  note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   268
  show "x \<le> y \<or> y \<le> x" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   269
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   270
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   271
text\<open>Constant Vectors\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   272
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   273
definition "vec x = (\<chi> i. x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   274
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   275
text\<open>Also the scalar-vector multiplication.\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   276
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80768
diff changeset
   277
definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl \<open>*s\<close> 70)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   278
  where "c *s x = (\<chi> i. c * (x$i))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   279
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   280
text \<open>scalar product\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   281
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   282
definition scalar_product :: "'a :: semiring_1 ^ 'n \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a" where
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   283
  "scalar_product v w = (\<Sum> i \<in> UNIV. v $ i * w $ i)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   284
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   285
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   286
subsection \<open>Real vector space\<close>
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   287
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   288
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (real_vector, finite) real_vector
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   289
begin
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   290
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   291
definition\<^marker>\<open>tag important\<close> "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   292
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   293
lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   294
  unfolding scaleR_vec_def by simp
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   295
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   296
instance\<^marker>\<open>tag unimportant\<close>
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60420
diff changeset
   297
  by standard (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   298
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   299
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   300
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   301
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   302
subsection \<open>Topological space\<close>
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   303
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   304
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (topological_space, finite) topological_space
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   305
begin
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   306
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   307
definition\<^marker>\<open>tag important\<close> [code del]:
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   308
  "open (S :: ('a ^ 'b) set) \<longleftrightarrow>
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   309
    (\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and>
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   310
      (\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   311
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   312
instance\<^marker>\<open>tag unimportant\<close> proof
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   313
  show "open (UNIV :: ('a ^ 'b) set)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   314
    unfolding open_vec_def by auto
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   315
next
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   316
  fix S T :: "('a ^ 'b) set"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   317
  assume "open S" "open T" thus "open (S \<inter> T)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   318
    unfolding open_vec_def
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   319
    apply clarify
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   320
    apply (drule (1) bspec)+
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   321
    apply (clarify, rename_tac Sa Ta)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   322
    apply (rule_tac x="\<lambda>i. Sa i \<inter> Ta i" in exI)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   323
    apply (simp add: open_Int)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   324
    done
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   325
next
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   326
  fix K :: "('a ^ 'b) set set"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   327
  assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   328
    unfolding open_vec_def
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   329
    by (metis Union_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   330
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   331
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   332
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   333
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   334
lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   335
  unfolding open_vec_def by auto
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   336
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   337
lemma open_vimage_vec_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   338
  unfolding open_vec_def
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   339
  apply clarify
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   340
  apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp)
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   341
  done
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   342
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   343
lemma closed_vimage_vec_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   344
  unfolding closed_open vimage_Compl [symmetric]
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   345
  by (rule open_vimage_vec_nth)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   346
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   347
lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   348
proof -
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   349
  have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   350
  thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   351
    by (simp add: closed_INT closed_vimage_vec_nth)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   352
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   353
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   354
lemma tendsto_vec_nth [tendsto_intros]:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   355
  assumes "((\<lambda>x. f x) \<longlongrightarrow> a) net"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   356
  shows "((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   357
proof (rule topological_tendstoI)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   358
  fix S assume "open S" "a $ i \<in> S"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   359
  then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   360
    by (simp_all add: open_vimage_vec_nth)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   361
  with assms have "eventually (\<lambda>x. f x \<in> (\<lambda>y. y $ i) -` S) net"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   362
    by (rule topological_tendstoD)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   363
  then show "eventually (\<lambda>x. f x $ i \<in> S) net"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   364
    by simp
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   365
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   366
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   367
lemma isCont_vec_nth [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x $ i) a"
44631
6820684c7a58 generalize lemma isCont_vec_nth
huffman
parents: 44630
diff changeset
   368
  unfolding isCont_def by (rule tendsto_vec_nth)
6820684c7a58 generalize lemma isCont_vec_nth
huffman
parents: 44630
diff changeset
   369
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   370
lemma vec_tendstoI:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   371
  assumes "\<And>i. ((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   372
  shows "((\<lambda>x. f x) \<longlongrightarrow> a) net"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   373
proof (rule topological_tendstoI)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   374
  fix S assume "open S" and "a \<in> S"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   375
  then obtain A where A: "\<And>i. open (A i)" "\<And>i. a $ i \<in> A i"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   376
    and S: "\<And>y. \<forall>i. y $ i \<in> A i \<Longrightarrow> y \<in> S"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   377
    unfolding open_vec_def by metis
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   378
  have "\<And>i. eventually (\<lambda>x. f x $ i \<in> A i) net"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   379
    using assms A by (rule topological_tendstoD)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   380
  hence "eventually (\<lambda>x. \<forall>i. f x $ i \<in> A i) net"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   381
    by (rule eventually_all_finite)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   382
  thus "eventually (\<lambda>x. f x \<in> S) net"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61169
diff changeset
   383
    by (rule eventually_mono, simp add: S)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   384
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   385
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   386
lemma tendsto_vec_lambda [tendsto_intros]:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   387
  assumes "\<And>i. ((\<lambda>x. f x i) \<longlongrightarrow> a i) net"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   388
  shows "((\<lambda>x. \<chi> i. f x i) \<longlongrightarrow> (\<chi> i. a i)) net"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   389
  using assms by (simp add: vec_tendstoI)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   390
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   391
lemma open_image_vec_nth: assumes "open S" shows "open ((\<lambda>x. x $ i) ` S)"
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   392
proof (rule openI)
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   393
  fix a assume "a \<in> (\<lambda>x. x $ i) ` S"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   394
  then obtain z where "a = z $ i" and "z \<in> S" ..
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   395
  then obtain A where A: "\<forall>i. open (A i) \<and> z $ i \<in> A i"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   396
    and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   397
    using \<open>open S\<close> unfolding open_vec_def by auto
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   398
  hence "A i \<subseteq> (\<lambda>x. x $ i) ` S"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   399
    by (clarsimp, rule_tac x="\<chi> j. if j = i then x else z $ j" in image_eqI,
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   400
      simp_all)
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   401
  hence "open (A i) \<and> a \<in> A i \<and> A i \<subseteq> (\<lambda>x. x $ i) ` S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   402
    using A \<open>a = z $ i\<close> by simp
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   403
  then show "\<exists>T. open T \<and> a \<in> T \<and> T \<subseteq> (\<lambda>x. x $ i) ` S" by - (rule exI)
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   404
qed
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   405
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   406
instance\<^marker>\<open>tag unimportant\<close> vec :: (perfect_space, finite) perfect_space
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   407
proof
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   408
  fix x :: "'a ^ 'b" show "\<not> open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   409
  proof
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   410
    assume "open {x}"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   411
    hence "\<forall>i. open ((\<lambda>x. x $ i) ` {x})" by (fast intro: open_image_vec_nth)
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   412
    hence "\<forall>i. open {x $ i}" by simp
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   413
    thus "False" by (simp add: not_open_singleton)
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   414
  qed
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   415
qed
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   416
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44282
diff changeset
   417
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   418
subsection \<open>Metric space\<close>
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   419
(* TODO: Product of uniform spaces and compatibility with metric_spaces! *)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   420
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   421
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) dist
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   422
begin
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   423
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   424
definition\<^marker>\<open>tag important\<close>
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   425
  "dist x y = L2_set (\<lambda>i. dist (x$i) (y$i)) UNIV"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   426
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   427
instance ..
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   428
end
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   429
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   430
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) uniformity_dist
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   431
begin
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   432
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   433
definition\<^marker>\<open>tag important\<close> [code del]:
67731
184c293f0a33 clarified use of vec type syntax;
wenzelm
parents: 67686
diff changeset
   434
  "(uniformity :: (('a^'b::_) \<times> ('a^'b::_)) filter) =
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
   435
    (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   436
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   437
instance\<^marker>\<open>tag unimportant\<close>
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   438
  by standard (rule uniformity_vec_def)
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   439
end
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   440
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   441
declare uniformity_Abort[where 'a="'a :: metric_space ^ 'b :: finite", code]
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   442
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   443
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) metric_space
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   444
begin
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   445
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   446
proposition dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   447
  unfolding dist_vec_def by (rule member_le_L2_set) simp_all
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   448
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   449
instance proof
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   450
  fix x y :: "'a ^ 'b"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   451
  show "dist x y = 0 \<longleftrightarrow> x = y"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   452
    unfolding dist_vec_def
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   453
    by (simp add: L2_set_eq_0_iff vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   454
next
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   455
  fix x y z :: "'a ^ 'b"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   456
  show "dist x y \<le> dist x z + dist y z"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   457
    unfolding dist_vec_def
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   458
    apply (rule order_trans [OF _ L2_set_triangle_ineq])
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   459
    apply (simp add: L2_set_mono dist_triangle2)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   460
    done
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   461
next
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   462
  fix S :: "('a ^ 'b) set"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   463
  have *: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
44630
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   464
  proof
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   465
    assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   466
    proof
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   467
      fix x assume "x \<in> S"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   468
      obtain A where A: "\<forall>i. open (A i)" "\<forall>i. x $ i \<in> A i"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   469
        and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   470
        using \<open>open S\<close> and \<open>x \<in> S\<close> unfolding open_vec_def by metis
44630
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   471
      have "\<forall>i\<in>UNIV. \<exists>r>0. \<forall>y. dist y (x $ i) < r \<longrightarrow> y \<in> A i"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   472
        using A unfolding open_dist by simp
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   473
      hence "\<exists>r. \<forall>i\<in>UNIV. 0 < r i \<and> (\<forall>y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i)"
44681
49ef76b4a634 remove duplicate lemma finite_choice in favor of finite_set_choice
huffman
parents: 44631
diff changeset
   474
        by (rule finite_set_choice [OF finite])
44630
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   475
      then obtain r where r1: "\<forall>i. 0 < r i"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   476
        and r2: "\<forall>i y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i" by fast
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   477
      have "0 < Min (range r) \<and> (\<forall>y. dist y x < Min (range r) \<longrightarrow> y \<in> S)"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   478
        by (simp add: r1 r2 S le_less_trans [OF dist_vec_nth_le])
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   479
      thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" ..
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   480
    qed
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   481
  next
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   482
    assume *: "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" show "open S"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   483
    proof (unfold open_vec_def, rule)
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   484
      fix x assume "x \<in> S"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   485
      then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   486
        using * by fast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62397
diff changeset
   487
      define r where [abs_def]: "r i = e / sqrt (of_nat CARD('b))" for i :: 'b
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   488
      from \<open>0 < e\<close> have r: "\<forall>i. 0 < r i"
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 54230
diff changeset
   489
        unfolding r_def by simp_all
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   490
      from \<open>0 < e\<close> have e: "e = L2_set r UNIV"
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   491
        unfolding r_def by (simp add: L2_set_constant)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62397
diff changeset
   492
      define A where "A i = {y. dist (x $ i) y < r i}" for i
44630
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   493
      have "\<forall>i. open (A i) \<and> x $ i \<in> A i"
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   494
        unfolding A_def by (simp add: open_ball r)
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   495
      moreover have "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S"
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   496
        by (simp add: A_def S dist_vec_def e L2_set_strict_mono dist_commute)
44630
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   497
      ultimately show "\<exists>A. (\<forall>i. open (A i) \<and> x $ i \<in> A i) \<and>
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   498
        (\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S)" by metis
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   499
    qed
d08cb39b628a convert proof to Isar-style
huffman
parents: 44571
diff changeset
   500
  qed
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   501
  show "open S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)"
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   502
    unfolding * eventually_uniformity_metric
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61973
diff changeset
   503
    by (simp del: split_paired_All add: dist_vec_def dist_commute)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   504
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   505
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   506
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   507
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   508
lemma Cauchy_vec_nth:
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   509
  "Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   510
  unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_vec_nth_le])
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   511
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   512
lemma vec_CauchyI:
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   513
  fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   514
  assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   515
  shows "Cauchy (\<lambda>n. X n)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   516
proof (rule metric_CauchyI)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   517
  fix r :: real assume "0 < r"
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 54230
diff changeset
   518
  hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62397
diff changeset
   519
  define N where "N i = (LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s)" for i
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62397
diff changeset
   520
  define M where "M = Max (range N)"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   521
  have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   522
    using X \<open>0 < ?s\<close> by (rule metric_CauchyD)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   523
  hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   524
    unfolding N_def by (rule LeastI_ex)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   525
  hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   526
    unfolding M_def by simp
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   527
  {
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   528
    fix m n :: nat
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   529
    assume "M \<le> m" "M \<le> n"
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   530
    have "dist (X m) (X n) = L2_set (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   531
      unfolding dist_vec_def ..
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   532
    also have "\<dots> \<le> sum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   533
      by (rule L2_set_le_sum [OF zero_le_dist])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   534
    also have "\<dots> < sum (\<lambda>i::'n. ?s) UNIV"
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   535
      by (rule sum_strict_mono, simp_all add: M \<open>M \<le> m\<close> \<open>M \<le> n\<close>)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   536
    also have "\<dots> = r"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   537
      by simp
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   538
    finally have "dist (X m) (X n) < r" .
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   539
  }
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   540
  hence "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   541
    by simp
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   542
  then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" ..
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   543
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   544
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   545
instance\<^marker>\<open>tag unimportant\<close> vec :: (complete_space, finite) complete_space
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   546
proof
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   547
  fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X"
61969
e01015e49041 more symbols;
wenzelm
parents: 61810
diff changeset
   548
  have "\<And>i. (\<lambda>n. X n $ i) \<longlonglongrightarrow> lim (\<lambda>n. X n $ i)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   549
    using Cauchy_vec_nth [OF \<open>Cauchy X\<close>]
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   550
    by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
61969
e01015e49041 more symbols;
wenzelm
parents: 61810
diff changeset
   551
  hence "X \<longlonglongrightarrow> vec_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   552
    by (simp add: vec_tendstoI)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   553
  then show "convergent X"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   554
    by (rule convergentI)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   555
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   556
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   557
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   558
subsection \<open>Normed vector space\<close>
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   559
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   560
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (real_normed_vector, finite) real_normed_vector
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   561
begin
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   562
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   563
definition\<^marker>\<open>tag important\<close> "norm x = L2_set (\<lambda>i. norm (x$i)) UNIV"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   564
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   565
definition\<^marker>\<open>tag important\<close> "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   566
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   567
instance\<^marker>\<open>tag unimportant\<close> proof
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   568
  fix a :: real and x y :: "'a ^ 'b"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   569
  show "norm x = 0 \<longleftrightarrow> x = 0"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   570
    unfolding norm_vec_def
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   571
    by (simp add: L2_set_eq_0_iff vec_eq_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   572
  show "norm (x + y) \<le> norm x + norm y"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   573
    unfolding norm_vec_def
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   574
    apply (rule order_trans [OF _ L2_set_triangle_ineq])
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   575
    apply (simp add: L2_set_mono norm_triangle_ineq)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   576
    done
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   577
  show "norm (scaleR a x) = \<bar>a\<bar> * norm x"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   578
    unfolding norm_vec_def
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   579
    by (simp add: L2_set_right_distrib)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   580
  show "sgn x = scaleR (inverse (norm x)) x"
44141
0697c01ff3ea follow standard naming scheme for sgn_vec_def
huffman
parents: 44136
diff changeset
   581
    by (rule sgn_vec_def)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   582
  show "dist x y = norm (x - y)"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   583
    unfolding dist_vec_def norm_vec_def
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   584
    by (simp add: dist_norm)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   585
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   586
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   587
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   588
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   589
lemma norm_nth_le: "norm (x $ i) \<le> norm x"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   590
unfolding norm_vec_def
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   591
by (rule member_le_L2_set) simp_all
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   592
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   593
lemma norm_le_componentwise_cart:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   594
  fixes x :: "'a::real_normed_vector^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   595
  assumes "\<And>i. norm(x$i) \<le> norm(y$i)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   596
  shows "norm x \<le> norm y"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   597
  unfolding norm_vec_def
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   598
  by (rule L2_set_mono) (auto simp: assms)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   599
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   600
lemma component_le_norm_cart: "\<bar>x$i\<bar> \<le> norm x"
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   601
  by (metis norm_nth_le real_norm_def)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   602
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   603
lemma norm_bound_component_le_cart: "norm x \<le> e ==> \<bar>x$i\<bar> \<le> e"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   604
  by (metis component_le_norm_cart order_trans)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   605
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   606
lemma norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   607
  by (metis component_le_norm_cart le_less_trans)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   608
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   609
lemma norm_le_l1_cart: "norm x \<le> sum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   610
  by (simp add: norm_vec_def L2_set_le_sum)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   611
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   612
lemma bounded_linear_vec_nth[intro]: "bounded_linear (\<lambda>x. x $ i)"
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   613
proof
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   614
  show "\<exists>K. \<forall>x. norm (x $ i) \<le> norm x * K"
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   615
    by (metis mult.commute mult.left_neutral norm_nth_le)
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   616
qed auto
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   617
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   618
instance vec :: (banach, finite) banach ..
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   619
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   620
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   621
subsection \<open>Inner product space\<close>
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   622
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   623
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (real_inner, finite) real_inner
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   624
begin
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   625
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   626
definition\<^marker>\<open>tag important\<close> "inner x y = sum (\<lambda>i. inner (x$i) (y$i)) UNIV"
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   627
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   628
instance\<^marker>\<open>tag unimportant\<close> proof
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   629
  fix r :: real and x y z :: "'a ^ 'b"
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   630
  show "inner x y = inner y x"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   631
    unfolding inner_vec_def
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   632
    by (simp add: inner_commute)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   633
  show "inner (x + y) z = inner x z + inner y z"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   634
    unfolding inner_vec_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   635
    by (simp add: inner_add_left sum.distrib)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   636
  show "inner (scaleR r x) y = r * inner x y"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   637
    unfolding inner_vec_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   638
    by (simp add: sum_distrib_left)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   639
  show "0 \<le> inner x x"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   640
    unfolding inner_vec_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   641
    by (simp add: sum_nonneg)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   642
  show "inner x x = 0 \<longleftrightarrow> x = 0"
44136
e63ad7d5158d more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents: 44135
diff changeset
   643
    unfolding inner_vec_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   644
    by (simp add: vec_eq_iff sum_nonneg_eq_0_iff)
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   645
  show "norm x = sqrt (inner x x)"
67155
9e5b05d54f9d canonical name
nipkow
parents: 66453
diff changeset
   646
    unfolding inner_vec_def norm_vec_def L2_set_def
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   647
    by (simp add: power2_norm_eq_inner)
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   648
qed
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   649
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   650
end
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   651
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   652
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   653
subsection \<open>Euclidean space\<close>
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   654
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
   655
text \<open>Vectors pointing along a single axis.\<close>
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   656
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   657
definition\<^marker>\<open>tag important\<close> "axis k x = (\<chi> i. if i = k then x else 0)"
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   658
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   659
lemma axis_nth [simp]: "axis i x $ i = x"
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   660
  unfolding axis_def by simp
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   661
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   662
lemma axis_eq_axis: "axis i x = axis j y \<longleftrightarrow> x = y \<and> i = j \<or> x = 0 \<and> y = 0"
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   663
  unfolding axis_def vec_eq_iff by auto
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   664
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   665
lemma inner_axis_axis:
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   666
  "inner (axis i x) (axis j y) = (if i = j then inner x y else 0)"
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
   667
  by (simp add: inner_vec_def axis_def sum.neutral sum.remove [of _ j])
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   668
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   669
lemma inner_axis: "inner x (axis i y) = inner (x $ i) y"
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   670
  by (simp add: inner_vec_def axis_def sum.remove [where x=i])
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   671
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   672
lemma inner_axis': "inner(axis i y) x = inner y (x $ i)"
67683
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   673
  by (simp add: inner_axis inner_commute)
817944aeac3f Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67673
diff changeset
   674
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   675
instantiation\<^marker>\<open>tag unimportant\<close> vec :: (euclidean_space, finite) euclidean_space
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   676
begin
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   677
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   678
definition\<^marker>\<open>tag important\<close> "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   679
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   680
instance\<^marker>\<open>tag unimportant\<close> proof
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   681
  show "(Basis :: ('a ^ 'b) set) \<noteq> {}"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   682
    unfolding Basis_vec_def by simp
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   683
next
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   684
  show "finite (Basis :: ('a ^ 'b) set)"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   685
    unfolding Basis_vec_def by simp
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   686
next
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   687
  fix u v :: "'a ^ 'b"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   688
  assume "u \<in> Basis" and "v \<in> Basis"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   689
  thus "inner u v = (if u = v then 1 else 0)"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   690
    unfolding Basis_vec_def
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   691
    by (auto simp add: inner_axis_axis axis_eq_axis inner_Basis)
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   692
next
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   693
  fix x :: "'a ^ 'b"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   694
  show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   695
    unfolding Basis_vec_def
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44165
diff changeset
   696
    by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50252
diff changeset
   697
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50252
diff changeset
   698
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   699
proposition DIM_cart [simp]: "DIM('a^'b) = CARD('b) * DIM('a)"
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   700
proof -
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   701
  have "card (\<Union>i::'b. \<Union>u::'a\<in>Basis. {axis i u}) = (\<Sum>i::'b\<in>UNIV. card (\<Union>u::'a\<in>Basis. {axis i u}))"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   702
    by (rule card_UN_disjoint) (auto simp: axis_eq_axis) 
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   703
  also have "... = CARD('b) * DIM('a)"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   704
    by (subst card_UN_disjoint) (auto simp: axis_eq_axis)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   705
  finally show ?thesis
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   706
    by (simp add: Basis_vec_def)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   707
qed
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   708
36591
df38e0c5c123 move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 36590
diff changeset
   709
end
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   710
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   711
lemma norm_axis_1 [simp]: "norm (axis m (1::real)) = 1"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   712
  by (simp add: inner_axis' norm_eq_1)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   713
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   714
lemma sum_norm_allsubsets_bound_cart:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   715
  fixes f:: "'a \<Rightarrow> real ^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   716
  assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (sum f Q) \<le> e"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   717
  shows "sum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) *  e"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   718
  using sum_norm_allsubsets_bound[OF assms]
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   719
  by simp
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   720
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   721
lemma cart_eq_inner_axis: "a $ i = inner a (axis i 1)"
62397
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62102
diff changeset
   722
  by (simp add: inner_axis)
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62102
diff changeset
   723
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   724
lemma axis_eq_0_iff [simp]:
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   725
  shows "axis m x = 0 \<longleftrightarrow> x = 0"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   726
  by (simp add: axis_def vec_eq_iff)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   727
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   728
lemma axis_in_Basis_iff [simp]: "axis i a \<in> Basis \<longleftrightarrow> a \<in> Basis"
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   729
  by (auto simp: Basis_vec_def axis_eq_axis)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   730
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   731
text\<open>Mapping each basis element to the corresponding finite index\<close>
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   732
definition axis_index :: "('a::comm_ring_1)^'n \<Rightarrow> 'n" where "axis_index v \<equiv> SOME i. v = axis i 1"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   733
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   734
lemma axis_inverse:
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   735
  fixes v :: "real^'n"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   736
  assumes "v \<in> Basis"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   737
  shows "\<exists>i. v = axis i 1"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   738
proof -
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   739
  have "v \<in> (\<Union>n. \<Union>r\<in>Basis. {axis n r})"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   740
    using assms Basis_vec_def by blast
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   741
  then show ?thesis
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   742
    by (force simp add: vec_eq_iff)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   743
qed
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   744
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   745
lemma axis_index:
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   746
  fixes v :: "real^'n"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   747
  assumes "v \<in> Basis"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   748
  shows "v = axis (axis_index v) 1"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   749
  by (metis (mono_tags) assms axis_inverse axis_index_def someI_ex)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   750
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   751
lemma axis_index_axis [simp]:
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   752
  fixes UU :: "real^'n"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   753
  shows "(axis_index (axis u 1 :: real^'n)) = (u::'n)"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   754
  by (simp add: axis_eq_axis axis_index_def)
62397
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62102
diff changeset
   755
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   756
subsection\<^marker>\<open>tag unimportant\<close> \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   757
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   758
lemma sum_cong_aux:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   759
  "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> sum f A = sum g A"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   760
  by (auto intro: sum.cong)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   761
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   762
hide_fact (open) sum_cong_aux
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   763
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   764
method_setup vector = \<open>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   765
let
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   766
  val ss1 =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
   767
    simpset_of (put_simpset HOL_basic_ss \<^context>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   768
      addsimps [@{thm sum.distrib} RS sym,
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   769
      @{thm sum_subtractf} RS sym, @{thm sum_distrib_left},
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   770
      @{thm sum_distrib_right}, @{thm sum_negf} RS sym])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   771
  val ss2 =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
   772
    simpset_of (\<^context> addsimps
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   773
             [@{thm plus_vec_def}, @{thm times_vec_def},
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   774
              @{thm minus_vec_def}, @{thm uminus_vec_def},
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   775
              @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def},
71174
nipkow
parents: 71044
diff changeset
   776
              @{thm scaleR_vec_def}, @{thm vector_scalar_mult_def}])
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   777
  fun vector_arith_tac ctxt ths =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   778
    simp_tac (put_simpset ss1 ctxt)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   779
    THEN' (fn i => resolve_tac ctxt @{thms Finite_Cartesian_Product.sum_cong_aux} i
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   780
         ORELSE resolve_tac ctxt @{thms sum.neutral} i
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   781
         ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   782
    (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   783
    THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   784
in
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   785
  Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths))
44135
18b4ab6854f1 move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
huffman
parents: 42290
diff changeset
   786
end
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   787
\<close> "lift trivial vector statements to real arith statements"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   788
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   789
lemma vec_0[simp]: "vec 0 = 0" by vector
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   790
lemma vec_1[simp]: "vec 1 = 1" by vector
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   791
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   792
lemma vec_inj[simp]: "vec x = vec y \<longleftrightarrow> x = y" by vector
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   793
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   794
lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   795
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   796
lemma vec_add: "vec(x + y) = vec x + vec y"  by vector
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   797
lemma vec_sub: "vec(x - y) = vec x - vec y" by vector
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   798
lemma vec_cmul: "vec(c * x) = c *s vec x " by vector
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   799
lemma vec_neg: "vec(- x) = - vec x " by vector
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   800
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   801
lemma vec_scaleR: "vec(c * x) = c *\<^sub>R vec x"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   802
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   803
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   804
lemma vec_sum:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   805
  assumes "finite S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   806
  shows "vec(sum f S) = sum (vec \<circ> f) S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   807
  using assms
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   808
proof induct
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   809
  case empty
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   810
  then show ?case by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   811
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   812
  case insert
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   813
  then show ?case by (auto simp add: vec_add)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   814
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   815
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   816
text\<open>Obvious "component-pushing".\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   817
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   818
lemma vec_component [simp]: "vec x $ i = x"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   819
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   820
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   821
lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   822
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   823
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   824
lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   825
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   826
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   827
lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   828
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   829
lemmas\<^marker>\<open>tag unimportant\<close> vector_component =
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   830
  vec_component vector_add_component vector_mult_component
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   831
  vector_smult_component vector_minus_component vector_uminus_component
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   832
  vector_scaleR_component cond_component
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   833
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   834
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   835
subsection\<^marker>\<open>tag unimportant\<close> \<open>Some frequently useful arithmetic lemmas over vectors\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   836
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   837
instance vec :: (semigroup_mult, finite) semigroup_mult
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   838
  by standard (vector mult.assoc)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   839
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   840
instance vec :: (monoid_mult, finite) monoid_mult
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   841
  by standard vector+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   842
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   843
instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   844
  by standard (vector mult.commute)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   845
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   846
instance vec :: (comm_monoid_mult, finite) comm_monoid_mult
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   847
  by standard vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   848
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   849
instance vec :: (semiring, finite) semiring
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   850
  by standard (vector field_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   851
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   852
instance vec :: (semiring_0, finite) semiring_0
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   853
  by standard (vector field_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   854
instance vec :: (semiring_1, finite) semiring_1
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   855
  by standard vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   856
instance vec :: (comm_semiring, finite) comm_semiring
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   857
  by standard (vector field_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   858
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   859
instance vec :: (comm_semiring_0, finite) comm_semiring_0 ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   860
instance vec :: (semiring_0_cancel, finite) semiring_0_cancel ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   861
instance vec :: (comm_semiring_0_cancel, finite) comm_semiring_0_cancel ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   862
instance vec :: (ring, finite) ring ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   863
instance vec :: (semiring_1_cancel, finite) semiring_1_cancel ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   864
instance vec :: (comm_semiring_1, finite) comm_semiring_1 ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   865
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   866
instance vec :: (ring_1, finite) ring_1 ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   867
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   868
instance vec :: (real_algebra, finite) real_algebra
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   869
  by standard (simp_all add: vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   870
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   871
instance vec :: (real_algebra_1, finite) real_algebra_1 ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   872
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   873
lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   874
proof (induct n)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   875
  case 0
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   876
  then show ?case by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   877
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   878
  case Suc
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   879
  then show ?case by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   880
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   881
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   882
lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   883
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   884
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   885
lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   886
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   887
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   888
instance vec :: (semiring_char_0, finite) semiring_char_0
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   889
proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   890
  fix m n :: nat
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   891
  show "inj (of_nat :: nat \<Rightarrow> 'a ^ 'b)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   892
    by (auto intro!: injI simp add: vec_eq_iff of_nat_index)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   893
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   894
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   895
instance vec :: (numeral, finite) numeral ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   896
instance vec :: (semiring_numeral, finite) semiring_numeral ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   897
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   898
lemma numeral_index [simp]: "numeral w $ i = numeral w"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   899
  by (induct w) (simp_all only: numeral.simps vector_add_component one_index)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   900
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   901
lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   902
  by (simp only: vector_uminus_component numeral_index)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   903
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   904
instance vec :: (comm_ring_1, finite) comm_ring_1 ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   905
instance vec :: (ring_char_0, finite) ring_char_0 ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   906
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   907
lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   908
  by (vector mult.assoc)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   909
lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   910
  by (vector field_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   911
lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   912
  by (vector field_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   913
lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   914
lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   915
lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   916
  by (vector field_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   917
lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   918
lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   919
lemma vector_sneg_minus1: "-x = (-1::'a::ring_1) *s x" by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   920
lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   921
lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   922
  by (vector field_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   923
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   924
lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   925
  by (simp add: vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   926
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   927
lemma Vector_Spaces_linear_vec [simp]: "Vector_Spaces.linear (*) vector_scalar_mult vec"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   928
  by unfold_locales (vector algebra_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   929
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   930
lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   931
  by vector
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   932
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   933
lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::'a::field) \<or> x = y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   934
  by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   935
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   936
lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::'a::field) = b \<or> x = 0"
71840
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71174
diff changeset
   937
  by (subst eq_iff_diff_eq_0, subst vector_sub_rdistrib [symmetric]) simp
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   938
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   939
lemma scalar_mult_eq_scaleR [abs_def]: "c *s x = c *\<^sub>R x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   940
  unfolding scaleR_vec_def vector_scalar_mult_def by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   941
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   942
lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   943
  unfolding dist_norm scalar_mult_eq_scaleR
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   944
  unfolding scaleR_right_diff_distrib[symmetric] by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   945
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   946
lemma sum_component [simp]:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   947
  fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   948
  shows "(sum f S)$i = sum (\<lambda>x. (f x)$i) S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   949
proof (cases "finite S")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   950
  case True
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   951
  then show ?thesis by induct simp_all
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   952
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   953
  case False
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   954
  then show ?thesis by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   955
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   956
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   957
lemma sum_eq: "sum f S = (\<chi> i. sum (\<lambda>x. (f x)$i ) S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   958
  by (simp add: vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   959
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   960
lemma sum_cmul:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   961
  fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   962
  shows "sum (\<lambda>x. c *s f x) S = c *s sum f S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   963
  by (simp add: vec_eq_iff sum_distrib_left)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   964
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   965
lemma linear_vec [simp]: "linear vec"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   966
  using Vector_Spaces_linear_vec
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   967
  by unfold_locales (vector algebra_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   968
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   969
subsection \<open>Matrix operations\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   970
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
   971
text\<open>Matrix notation. NB: an MxN matrix is of type \<^typ>\<open>'a^'n^'m\<close>, not \<^typ>\<open>'a^'m^'n\<close>\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   972
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   973
definition\<^marker>\<open>tag important\<close> map_matrix::"('a \<Rightarrow> 'b) \<Rightarrow> (('a, 'i::finite)vec, 'j::finite) vec \<Rightarrow> (('b, 'i)vec, 'j) vec" where
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   974
  "map_matrix f x = (\<chi> i j. f (x $ i $ j))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   975
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   976
lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   977
  by (simp add: map_matrix_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   978
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   979
definition\<^marker>\<open>tag important\<close> matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80768
diff changeset
   980
    (infixl \<open>**\<close> 70)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   981
  where "m ** m' == (\<chi> i j. sum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   982
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   983
definition\<^marker>\<open>tag important\<close> matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80768
diff changeset
   984
    (infixl \<open>*v\<close> 70)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   985
  where "m *v x \<equiv> (\<chi> i. sum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   986
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   987
definition\<^marker>\<open>tag important\<close> vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80768
diff changeset
   988
    (infixl \<open>v*\<close> 70)
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   989
  where "v v* m == (\<chi> j. sum (\<lambda>i. ((v$i) * (m$i)$j)) (UNIV :: 'm set)) :: 'a^'n"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   990
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   991
definition\<^marker>\<open>tag unimportant\<close> "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   992
definition\<^marker>\<open>tag unimportant\<close> transpose where
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   993
  "(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   994
definition\<^marker>\<open>tag unimportant\<close> "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   995
definition\<^marker>\<open>tag unimportant\<close> "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   996
definition\<^marker>\<open>tag unimportant\<close> "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
   997
definition\<^marker>\<open>tag unimportant\<close> "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   998
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
   999
lemma times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0" 
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1000
  by (simp add: matrix_matrix_mult_def zero_vec_def)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1001
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1002
lemma times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0" 
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1003
  by (simp add: matrix_matrix_mult_def zero_vec_def)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1004
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1005
lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1006
lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1007
  by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1008
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1009
lemma matrix_mul_lid [simp]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1010
  fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1011
  shows "mat 1 ** A = A"
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1012
  unfolding matrix_matrix_mult_def mat_def
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1013
  by (auto simp: if_distrib if_distribR sum.delta'[OF finite] cong: if_cong)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1014
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1015
lemma matrix_mul_rid [simp]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1016
  fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1017
  shows "A ** mat 1 = A"
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1018
  unfolding matrix_matrix_mult_def mat_def
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1019
  by (auto simp: if_distrib if_distribR sum.delta'[OF finite] cong: if_cong)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1020
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1021
lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1022
  apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc)
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1023
  using sum.swap by fastforce
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1024
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1025
lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1026
  apply (vector matrix_matrix_mult_def matrix_vector_mult_def
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1027
    sum_distrib_left sum_distrib_right mult.assoc)
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1028
  using sum.swap by fastforce
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1029
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1030
lemma vector_matrix_mul_assoc: "(x v* A) v* B = x v* (A**B)"
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1031
  apply (vector matrix_matrix_mult_def vector_matrix_mult_def
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1032
    sum_distrib_left sum_distrib_right mult.assoc)
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1033
  using sum.swap by fastforce
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1034
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1035
lemma scalar_matrix_assoc:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1036
  fixes A :: "('a::real_algebra_1)^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1037
  shows "k *\<^sub>R (A ** B) = (k *\<^sub>R A) ** B"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1038
  by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff scaleR_sum_right)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1039
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1040
lemma matrix_scalar_ac:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1041
  fixes A :: "('a::real_algebra_1)^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1042
  shows "A ** (k *\<^sub>R B) = k *\<^sub>R A ** B"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1043
  by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1044
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1045
lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1046
  apply (vector matrix_vector_mult_def mat_def)
71044
nipkow
parents: 70136
diff changeset
  1047
  apply (simp add: if_distrib if_distribR cong del: if_weak_cong)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1048
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1049
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1050
lemma vector_matrix_mul_rid [simp]:
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1051
  fixes v :: "('a::semiring_1)^'n"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1052
  shows "v v* mat 1 = v"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1053
  apply (vector vector_matrix_mult_def mat_def)
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1054
  apply (simp add: if_distrib if_distribR cong del: if_weak_cong)
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1055
  done
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1056
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1057
lemma matrix_transpose_mul:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1058
    "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1059
  by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1060
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1061
lemma matrix_mult_transpose_dot_column:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69668
diff changeset
  1062
  shows "transpose A ** A = (\<chi> i j. inner (column i A) (column j A))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69668
diff changeset
  1063
  by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69668
diff changeset
  1064
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1065
lemma matrix_mult_transpose_dot_row:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69668
diff changeset
  1066
  shows "A ** transpose A = (\<chi> i j. inner (row i A) (row j A))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69668
diff changeset
  1067
  by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69668
diff changeset
  1068
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1069
lemma matrix_eq:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1070
  fixes A B :: "'a::semiring_1 ^ 'n ^ 'm"
73976
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1071
  shows "A = B \<longleftrightarrow> (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1072
proof
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1073
  assume ?rhs
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1074
  then show ?lhs
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1075
    apply (subst vec_eq_iff)
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1076
    apply (clarsimp simp add: matrix_vector_mult_def if_distrib if_distribR vec_eq_iff cong: if_cong)
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1077
    apply (erule_tac x="axis ia 1" in allE)
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1078
    apply (erule_tac x="i" in allE)
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1079
    apply (auto simp add: if_distrib if_distribR axis_def
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1080
        sum.delta[OF finite] cong del: if_weak_cong)
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1081
    done
a5212df98387 simplified a few proofs
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1082
qed auto
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1083
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1084
lemma matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1085
  by (simp add: matrix_vector_mult_def inner_vec_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1086
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1087
lemma dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1088
  apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1089
  apply (subst sum.swap)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1090
  apply simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1091
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1092
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1093
lemma transpose_mat [simp]: "transpose (mat n) = mat n"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1094
  by (vector transpose_def mat_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1095
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1096
lemma transpose_transpose [simp]: "transpose(transpose A) = A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1097
  by (vector transpose_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1098
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1099
lemma row_transpose [simp]: "row i (transpose A) = column i A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1100
  by (simp add: row_def column_def transpose_def vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1101
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1102
lemma column_transpose [simp]: "column i (transpose A) = row i A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1103
  by (simp add: row_def column_def transpose_def vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1104
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1105
lemma rows_transpose [simp]: "rows(transpose A) = columns A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1106
  by (auto simp add: rows_def columns_def intro: set_eqI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1107
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1108
lemma columns_transpose [simp]: "columns(transpose A) = rows A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1109
  by (metis transpose_transpose rows_transpose)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1110
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1111
lemma transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1112
  unfolding transpose_def
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1113
  by (simp add: vec_eq_iff)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1114
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1115
lemma transpose_iff [iff]: "transpose A = transpose B \<longleftrightarrow> A = B"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1116
  by (metis transpose_transpose)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1117
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1118
lemma matrix_mult_sum:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1119
  "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1120
  by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1121
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1122
lemma vector_componentwise:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1123
  "(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1124
  by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1125
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1126
lemma basis_expansion: "sum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1127
  by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1128
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1129
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1130
text\<open>Correspondence between matrices and linear operators.\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1131
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
  1132
definition\<^marker>\<open>tag important\<close> matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1133
  where "matrix f = (\<chi> i j. (f(axis j 1))$i)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1134
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1135
lemma matrix_id_mat_1: "matrix id = mat 1"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1136
  by (simp add: mat_def matrix_def axis_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1137
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1138
lemma matrix_scaleR: "(matrix ((*\<^sub>R) r)) = mat r"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1139
  by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1140
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1141
lemma matrix_vector_mul_linear[intro, simp]: "linear (\<lambda>x. A *v (x::'a::real_algebra_1 ^ _))"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1142
  by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff field_simps sum_distrib_left
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1143
      sum.distrib scaleR_right.sum)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1144
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1145
lemma vector_matrix_left_distrib [algebra_simps]:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1146
  shows "(x + y) v* A = x v* A + y v* A"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1147
  unfolding vector_matrix_mult_def
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1148
  by (simp add: algebra_simps sum.distrib vec_eq_iff)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1149
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1150
lemma vector_matrix_mult_diff_distrib [algebra_simps]:
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1151
  fixes A :: "'a::ring_1^'n^'m"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1152
  shows "(x - y) v* A = x v* A - y v* A"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1153
  by (vector vector_matrix_mult_def sum_subtractf left_diff_distrib)
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1154
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1155
lemma matrix_vector_right_distrib [algebra_simps]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1156
  "A *v (x + y) = A *v x + A *v y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1157
  by (vector matrix_vector_mult_def sum.distrib distrib_left)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1158
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1159
lemma matrix_vector_mult_diff_distrib [algebra_simps]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1160
  fixes A :: "'a::ring_1^'n^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1161
  shows "A *v (x - y) = A *v x - A *v y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1162
  by (vector matrix_vector_mult_def sum_subtractf right_diff_distrib)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1163
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1164
lemma matrix_vector_mult_scaleR[algebra_simps]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1165
  fixes A :: "real^'n^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1166
  shows "A *v (c *\<^sub>R x) = c *\<^sub>R (A *v x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1167
  using linear_iff matrix_vector_mul_linear by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1168
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1169
lemma matrix_vector_mult_0_right [simp]: "A *v 0 = 0"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1170
  by (simp add: matrix_vector_mult_def vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1171
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1172
lemma matrix_vector_mult_0 [simp]: "0 *v w = 0"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1173
  by (simp add: matrix_vector_mult_def vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1174
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1175
lemma matrix_vector_mult_add_rdistrib [algebra_simps]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1176
  "(A + B) *v x = (A *v x) + (B *v x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1177
  by (vector matrix_vector_mult_def sum.distrib distrib_right)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1178
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1179
lemma matrix_vector_mult_diff_rdistrib [algebra_simps]:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1180
  fixes A :: "'a :: ring_1^'n^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1181
  shows "(A - B) *v x = (A *v x) - (B *v x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1182
  by (vector matrix_vector_mult_def sum_subtractf left_diff_distrib)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1183
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1184
lemma vector_matrix_mult_add_rdistrib [algebra_simps]:
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1185
  "x v* (A + B) = (x v* A) + (x v* B)"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1186
  by (vector vector_matrix_mult_def sum.distrib distrib_left)
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1187
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1188
lemma  vector_matrix_mult_diff_rdistrib [algebra_simps]:
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1189
  fixes A :: "'a :: ring_1^'n^'m"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1190
  shows "x v* (A - B) = (x v* A) - (x v* B)"
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1191
  by (vector vector_matrix_mult_def sum_subtractf right_diff_distrib)
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1192
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1193
lemma matrix_vector_column:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1194
  "(A::'a::comm_semiring_1^'n^_) *v x = sum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1195
  by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1196
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
  1197
subsection\<open>Inverse matrices  (not necessarily square)\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1198
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
  1199
definition\<^marker>\<open>tag important\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1200
  "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1201
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69720
diff changeset
  1202
definition\<^marker>\<open>tag important\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1203
  "matrix_inv(A:: 'a::semiring_1^'n^'m) =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1204
    (SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1205
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1206
lemma inj_matrix_vector_mult:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1207
  fixes A::"'a::field^'n^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1208
  assumes "invertible A"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1209
  shows "inj ((*v) A)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1210
  by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1211
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1212
lemma scalar_invertible:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1213
  fixes A :: "('a::real_algebra_1)^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1214
  assumes "k \<noteq> 0" and "invertible A"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1215
  shows "invertible (k *\<^sub>R A)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1216
proof -
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1217
  obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1218
    using assms unfolding invertible_def by auto
69272
15e9ed5b28fb isabelle update_cartouches -t;
wenzelm
parents: 69260
diff changeset
  1219
  with \<open>k \<noteq> 0\<close>
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1220
  have "(k *\<^sub>R A) ** ((1/k) *\<^sub>R A') = mat 1" "((1/k) *\<^sub>R A') ** (k *\<^sub>R A) = mat 1"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1221
    by (simp_all add: assms matrix_scalar_ac)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1222
  thus "invertible (k *\<^sub>R A)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1223
    unfolding invertible_def by auto
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1224
qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1225
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1226
lemma scalar_invertible_iff:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1227
  fixes A :: "('a::real_algebra_1)^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1228
  assumes "k \<noteq> 0" and "invertible A"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1229
  shows "invertible (k *\<^sub>R A) \<longleftrightarrow> k \<noteq> 0 \<and> invertible A"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1230
  by (simp add: assms scalar_invertible)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1231
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1232
lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v (x:: 'a::{comm_semiring_1}^'n)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1233
  unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1234
  by (simp add: mult.commute)
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1235
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1236
lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* (A:: 'a::{comm_semiring_1}^'m^'n)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1237
  unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
81471
1b24a570bcf7 patch to vector_matrix_mult by Alexander Pach
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
  1238
  by (simp add: mult.commute)
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1239
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1240
lemma vector_scalar_commute:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1241
  fixes A :: "'a::{field}^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1242
  shows "A *v (c *s x) = c *s (A *v x)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1243
  by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1244
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1245
lemma scalar_vector_matrix_assoc:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1246
  fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1247
  shows "(k *s x) v* A = k *s (x v* A)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1248
  by (metis transpose_matrix_vector vector_scalar_commute)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1249
 
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1250
lemma vector_matrix_mult_0 [simp]: "0 v* A = 0"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1251
  unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1252
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1253
lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1254
  unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1255
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1256
lemma scaleR_vector_matrix_assoc:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1257
  fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1258
  shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1259
  by (metis matrix_vector_mult_scaleR transpose_matrix_vector)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1260
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1261
lemma vector_scaleR_matrix_ac:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1262
  fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1263
  shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69684
diff changeset
  1264
proof -
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1265
  have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A"
81472
e43bff789ac0 Patch by Stepan Holub, plus tweaks
paulson <lp15@cam.ac.uk>
parents: 81471
diff changeset
  1266
    by (simp add: vector_matrix_mult_def algebra_simps)
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1267
  with scaleR_vector_matrix_assoc
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1268
  show "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1269
    by auto
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1270
qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1271
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  1272
end