| author | nipkow | 
| Thu, 13 Apr 2023 15:36:07 +1000 | |
| changeset 77830 | 0f2baf04b782 | 
| parent 73976 | a5212df98387 | 
| child 80768 | c7723cc15de8 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Finite_Cartesian_Product.thy | 
| 35253 | 2 | Author: Amine Chaieb, University of Cambridge | 
| 33175 | 3 | *) | 
| 4 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 5 | section \<open>Definition of Finite Cartesian Product Type\<close> | 
| 33175 | 6 | |
| 7 | theory Finite_Cartesian_Product | |
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39302diff
changeset | 8 | imports | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 9 | Euclidean_Space | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39302diff
changeset | 10 | L2_Norm | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66281diff
changeset | 11 | "HOL-Library.Numeral_Type" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66281diff
changeset | 12 | "HOL-Library.Countable_Set" | 
| 68188 
2af1f142f855
move FuncSet back to HOL-Library (amending 493b818e8e10)
 immler parents: 
68074diff
changeset | 13 | "HOL-Library.FuncSet" | 
| 33175 | 14 | begin | 
| 15 | ||
| 70136 | 16 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Finite Cartesian products, with indexing and lambdas\<close> | 
| 33175 | 17 | |
| 67731 | 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: 
44135diff
changeset | 19 | morphisms vec_nth vec_lambda .. | 
| 33175 | 20 | |
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 21 | declare vec_lambda_inject [simplified, simp] | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 22 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 23 | bundle vec_syntax begin | 
| 35254 
0f17eda72e60
binder notation for default print_mode -- to avoid strange output if "xsymbols" is not active;
 wenzelm parents: 
35253diff
changeset | 24 | notation | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 25 | vec_nth (infixl "$" 90) and | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 26 | vec_lambda (binder "\<chi>" 10) | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 27 | end | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 28 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 29 | bundle no_vec_syntax begin | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 30 | no_notation | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 31 | vec_nth (infixl "$" 90) and | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 32 | vec_lambda (binder "\<chi>" 10) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 33 | end | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 34 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 35 | unbundle vec_syntax | 
| 33175 | 36 | |
| 67731 | 37 | text \<open> | 
| 38 |   Concrete syntax for \<open>('a, 'b) vec\<close>:
 | |
| 39 |     \<^item> \<open>'a^'b\<close> becomes \<open>('a, 'b::finite) vec\<close>
 | |
| 40 |     \<^item> \<open>'a^'b::_\<close> becomes \<open>('a, 'b) vec\<close> without extra sort-constraint
 | |
| 41 | \<close> | |
| 67732 | 42 | syntax "_vec_type" :: "type \<Rightarrow> type \<Rightarrow> type" (infixl "^" 15) | 
| 60420 | 43 | parse_translation \<open> | 
| 52143 | 44 | let | 
| 69597 | 45 | fun vec t u = Syntax.const \<^type_syntax>\<open>vec\<close> $ t $ u; | 
| 52143 | 46 | fun finite_vec_tr [t, u] = | 
| 47 | (case Term_Position.strip_positions u of | |
| 48 | v as Free (x, _) => | |
| 49 | if Lexicon.is_tid x then | |
| 69597 | 50 | vec t (Syntax.const \<^syntax_const>\<open>_ofsort\<close> $ v $ | 
| 51 | Syntax.const \<^class_syntax>\<open>finite\<close>) | |
| 52143 | 52 | else vec t u | 
| 53 | | _ => vec t u) | |
| 54 | in | |
| 69597 | 55 | [(\<^syntax_const>\<open>_vec_type\<close>, K finite_vec_tr)] | 
| 52143 | 56 | end | 
| 60420 | 57 | \<close> | 
| 34290 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 58 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 59 | 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: 
44135diff
changeset | 60 | by (simp add: vec_nth_inject [symmetric] fun_eq_iff) | 
| 33175 | 61 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 62 | 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: 
44135diff
changeset | 63 | by (simp add: vec_lambda_inverse) | 
| 33175 | 64 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 65 | 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: 
44135diff
changeset | 66 | by (auto simp add: vec_eq_iff) | 
| 33175 | 67 | |
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 68 | 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: 
44135diff
changeset | 69 | by (simp add: vec_eq_iff) | 
| 33175 | 70 | |
| 69683 | 71 | subsection \<open>Cardinality of vectors\<close> | 
| 66281 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 72 | |
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 73 | instance vec :: (finite, finite) finite | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 74 | proof | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 75 |   show "finite (UNIV :: ('a, 'b) vec set)"
 | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 76 | proof (subst bij_betw_finite) | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 77 | 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: 
64267diff
changeset | 78 | 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: 
64267diff
changeset | 79 | have "finite (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 80 | by (intro finite_PiE) auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 81 | 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: 
64267diff
changeset | 82 | by auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 83 | finally show "finite \<dots>" . | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 84 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 85 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 86 | |
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 87 | lemma countable_PiE: | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 88 | "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: 
64267diff
changeset | 89 | 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: 
64267diff
changeset | 90 | |
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 91 | instance vec :: (countable, finite) countable | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 92 | proof | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 93 |   have "countable (UNIV :: ('a, 'b) vec set)"
 | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 94 | proof (rule countableI_bij2) | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 95 | 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: 
64267diff
changeset | 96 | 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: 
64267diff
changeset | 97 | have "countable (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 98 | by (intro countable_PiE) auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 99 | 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: 
64267diff
changeset | 100 | by auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 101 | finally show "countable \<dots>" . | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 102 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 103 |   thus "\<exists>t::('a, 'b) vec \<Rightarrow> nat. inj t"
 | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 104 | by (auto elim!: countableE) | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 105 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 106 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 107 | lemma infinite_UNIV_vec: | 
| 66281 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 108 | assumes "infinite (UNIV :: 'a set)" | 
| 67731 | 109 |   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: 
69684diff
changeset | 110 | proof (subst bij_betw_finite) | 
| 66281 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 111 | 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: 
64267diff
changeset | 112 | 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: 
64267diff
changeset | 113 | 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: 
64267diff
changeset | 114 | proof | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 115 | assume "finite ?A" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 116 | hence "finite ((\<lambda>f. f undefined) ` ?A)" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 117 | by (rule finite_imageI) | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 118 | also have "(\<lambda>f. f undefined) ` ?A = UNIV" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 119 | by auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 120 | finally show False | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 121 | using \<open>infinite (UNIV :: 'a set)\<close> by contradiction | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 122 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 123 | also have "?A = Pi UNIV (\<lambda>_. UNIV)" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 124 | by auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 125 | finally show "infinite (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" . | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 126 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 127 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 128 | proposition CARD_vec [simp]: | 
| 67731 | 129 |   "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: 
69684diff
changeset | 130 | proof (cases "finite (UNIV :: 'a set)") | 
| 66281 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 131 | case True | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 132 | show ?thesis | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 133 | proof (subst bij_betw_same_card) | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 134 | 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: 
64267diff
changeset | 135 | 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: 
64267diff
changeset | 136 |     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: 
64267diff
changeset | 137 | (is "_ = card ?A") | 
| 71044 | 138 | by (subst card_PiE) (auto) | 
| 66281 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 139 | also have "?A = Pi UNIV (\<lambda>_. UNIV)" | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 140 | by auto | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 141 |     finally show "card \<dots> = CARD('a) ^ CARD('b)" ..
 | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 142 | qed | 
| 
6ad54b84ca5d
facts about cardinality of vector type
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 143 | qed (simp_all add: infinite_UNIV_vec) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 144 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 145 | lemma countable_vector: | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 146 | fixes B:: "'n::finite \<Rightarrow> 'a set" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 147 | assumes "\<And>i. countable (B i)" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 148 |   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: 
67673diff
changeset | 149 | proof - | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 150 |   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: 
67673diff
changeset | 151 | proof - | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 152 | 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: 
67673diff
changeset | 153 | by (metis that PiE_iff UNIV_I vec_lambda_inverse) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 154 |     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: 
67673diff
changeset | 155 | by blast | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 156 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 157 |   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: 
67673diff
changeset | 158 | by blast | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 159 |   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: 
67673diff
changeset | 160 | by (metis finite_class.finite_UNIV countable_PiE assms) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 161 |   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: 
67673diff
changeset | 162 | by auto | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 163 | then show ?thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 164 | by (simp add: image_comp o_def vec_nth_inverse) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 165 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 166 | |
| 70136 | 167 | 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: 
36590diff
changeset | 168 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 169 | instantiation vec :: (zero, finite) zero | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 170 | begin | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 171 | definition "0 \<equiv> (\<chi> i. 0)" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 172 | instance .. | 
| 33175 | 173 | end | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 174 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 175 | instantiation vec :: (plus, finite) plus | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 176 | begin | 
| 67399 | 177 | 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: 
36590diff
changeset | 178 | instance .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 179 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 180 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 181 | instantiation vec :: (minus, finite) minus | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 182 | begin | 
| 67399 | 183 | 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: 
36590diff
changeset | 184 | instance .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 185 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 186 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 187 | instantiation vec :: (uminus, finite) uminus | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 188 | begin | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 189 | 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: 
36590diff
changeset | 190 | instance .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 191 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 192 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 193 | lemma zero_index [simp]: "0 $ i = 0" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 194 | unfolding zero_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 195 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 196 | 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: 
44135diff
changeset | 197 | unfolding plus_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 198 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 199 | 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: 
44135diff
changeset | 200 | unfolding minus_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 201 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 202 | 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: 
44135diff
changeset | 203 | unfolding uminus_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 204 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 205 | instance vec :: (semigroup_add, finite) semigroup_add | 
| 61169 | 206 | 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: 
36590diff
changeset | 207 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 208 | instance vec :: (ab_semigroup_add, finite) ab_semigroup_add | 
| 61169 | 209 | 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: 
36590diff
changeset | 210 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 211 | instance vec :: (monoid_add, finite) monoid_add | 
| 61169 | 212 | by standard (simp_all add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 213 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 214 | instance vec :: (comm_monoid_add, finite) comm_monoid_add | 
| 61169 | 215 | by standard (simp add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 216 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 217 | instance vec :: (cancel_semigroup_add, finite) cancel_semigroup_add | 
| 61169 | 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: 
36590diff
changeset | 219 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 220 | instance vec :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add | 
| 61169 | 221 | 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: 
36590diff
changeset | 222 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 223 | 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: 
36590diff
changeset | 224 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 225 | instance vec :: (group_add, finite) group_add | 
| 61169 | 226 | by standard (simp_all add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 227 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 228 | instance vec :: (ab_group_add, finite) ab_group_add | 
| 61169 | 229 | by standard (simp_all add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 230 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 231 | |
| 70136 | 232 | 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: 
67982diff
changeset | 233 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 234 | instantiation vec :: (times, finite) times | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 235 | begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 236 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 237 | 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: 
67982diff
changeset | 238 | instance .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 239 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 240 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 241 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 242 | instantiation vec :: (one, finite) one | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 243 | begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 244 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 245 | definition "1 \<equiv> (\<chi> i. 1)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 246 | instance .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 247 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 248 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 249 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 250 | instantiation vec :: (ord, finite) ord | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 251 | begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 252 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 253 | 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: 
67982diff
changeset | 254 | 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: 
67982diff
changeset | 255 | instance .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 256 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 257 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 258 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 259 | 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: 
67982diff
changeset | 260 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 261 | instance vec:: (order, finite) order | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 262 | 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: 
67982diff
changeset | 263 | 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: 
67982diff
changeset | 264 | |
| 69680 | 265 | 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: 
67982diff
changeset | 266 | proof | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 267 | 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: 
67982diff
changeset | 268 | proof - | 
| 69680 | 269 |     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: 
67982diff
changeset | 270 |     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: 
67982diff
changeset | 271 | 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: 
67982diff
changeset | 272 | 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: 
67982diff
changeset | 273 | qed | 
| 69680 | 274 | 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: 
67982diff
changeset | 275 | 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: 
67982diff
changeset | 276 | 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: 
67982diff
changeset | 277 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 278 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 279 | text\<open>Constant Vectors\<close> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 280 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 281 | definition "vec x = (\<chi> i. x)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 282 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 283 | 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: 
67982diff
changeset | 284 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 285 | definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 286 | 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: 
67982diff
changeset | 287 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 288 | text \<open>scalar product\<close> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 289 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 290 | 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: 
67982diff
changeset | 291 | "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: 
67982diff
changeset | 292 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 293 | |
| 69683 | 294 | subsection \<open>Real vector space\<close> | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 295 | |
| 70136 | 296 | 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: 
36590diff
changeset | 297 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 298 | |
| 70136 | 299 | 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: 
36590diff
changeset | 300 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 301 | 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: 
44135diff
changeset | 302 | unfolding scaleR_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 303 | |
| 70136 | 304 | instance\<^marker>\<open>tag unimportant\<close> | 
| 61169 | 305 | 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: 
36590diff
changeset | 306 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 307 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 308 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 309 | |
| 69683 | 310 | subsection \<open>Topological space\<close> | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 311 | |
| 70136 | 312 | 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: 
36590diff
changeset | 313 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 314 | |
| 70136 | 315 | definition\<^marker>\<open>tag important\<close> [code del]: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 316 |   "open (S :: ('a ^ 'b) set) \<longleftrightarrow>
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 317 | (\<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: 
36590diff
changeset | 318 | (\<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: 
36590diff
changeset | 319 | |
| 70136 | 320 | instance\<^marker>\<open>tag unimportant\<close> proof | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 321 |   show "open (UNIV :: ('a ^ 'b) set)"
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 322 | unfolding open_vec_def by auto | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 323 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 324 |   fix S T :: "('a ^ 'b) set"
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 325 | 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: 
44135diff
changeset | 326 | unfolding open_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 327 | apply clarify | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 328 | apply (drule (1) bspec)+ | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 329 | apply (clarify, rename_tac Sa Ta) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 330 | 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: 
36590diff
changeset | 331 | apply (simp add: open_Int) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 332 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 333 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 334 |   fix K :: "('a ^ 'b) set set"
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 335 | 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: 
44135diff
changeset | 336 | unfolding open_vec_def | 
| 73976 | 337 | by (metis Union_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 338 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 339 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 340 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 341 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 342 | 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: 
44135diff
changeset | 343 | unfolding open_vec_def by auto | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 344 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 345 | 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: 
44135diff
changeset | 346 | unfolding open_vec_def | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 347 | apply clarify | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 348 | 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: 
44135diff
changeset | 349 | done | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 350 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 351 | 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: 
44135diff
changeset | 352 | unfolding closed_open vimage_Compl [symmetric] | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 353 | by (rule open_vimage_vec_nth) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 354 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 355 | 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: 
36590diff
changeset | 356 | proof - | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 357 |   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: 
36590diff
changeset | 358 |   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: 
44135diff
changeset | 359 | 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: 
36590diff
changeset | 360 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 361 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 362 | lemma tendsto_vec_nth [tendsto_intros]: | 
| 61973 | 363 | assumes "((\<lambda>x. f x) \<longlongrightarrow> a) net" | 
| 364 | 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: 
69684diff
changeset | 365 | proof (rule topological_tendstoI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 366 | fix S assume "open S" "a $ i \<in> S" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 367 | 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: 
44135diff
changeset | 368 | by (simp_all add: open_vimage_vec_nth) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 369 | 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: 
36590diff
changeset | 370 | by (rule topological_tendstoD) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 371 | 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: 
36590diff
changeset | 372 | by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 373 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 374 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 375 | lemma isCont_vec_nth [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x $ i) a" | 
| 44631 | 376 | unfolding isCont_def by (rule tendsto_vec_nth) | 
| 377 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 378 | lemma vec_tendstoI: | 
| 61973 | 379 | assumes "\<And>i. ((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net" | 
| 380 | 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: 
69684diff
changeset | 381 | proof (rule topological_tendstoI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 382 | fix S assume "open S" and "a \<in> S" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 383 | 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: 
36590diff
changeset | 384 | 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: 
44135diff
changeset | 385 | unfolding open_vec_def by metis | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 386 | 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: 
36590diff
changeset | 387 | using assms A by (rule topological_tendstoD) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 388 | 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: 
36590diff
changeset | 389 | by (rule eventually_all_finite) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 390 | thus "eventually (\<lambda>x. f x \<in> S) net" | 
| 61810 | 391 | by (rule eventually_mono, simp add: S) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 392 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 393 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 394 | lemma tendsto_vec_lambda [tendsto_intros]: | 
| 61973 | 395 | assumes "\<And>i. ((\<lambda>x. f x i) \<longlongrightarrow> a i) net" | 
| 396 | 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: 
44135diff
changeset | 397 | using assms by (simp add: vec_tendstoI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 398 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 399 | 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: 
69684diff
changeset | 400 | proof (rule openI) | 
| 44571 | 401 | fix a assume "a \<in> (\<lambda>x. x $ i) ` S" | 
| 402 | then obtain z where "a = z $ i" and "z \<in> S" .. | |
| 403 | then obtain A where A: "\<forall>i. open (A i) \<and> z $ i \<in> A i" | |
| 404 | and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" | |
| 60420 | 405 | using \<open>open S\<close> unfolding open_vec_def by auto | 
| 44571 | 406 | hence "A i \<subseteq> (\<lambda>x. x $ i) ` S" | 
| 407 | by (clarsimp, rule_tac x="\<chi> j. if j = i then x else z $ j" in image_eqI, | |
| 408 | simp_all) | |
| 409 | hence "open (A i) \<and> a \<in> A i \<and> A i \<subseteq> (\<lambda>x. x $ i) ` S" | |
| 60420 | 410 | using A \<open>a = z $ i\<close> by simp | 
| 44571 | 411 | then show "\<exists>T. open T \<and> a \<in> T \<and> T \<subseteq> (\<lambda>x. x $ i) ` S" by - (rule exI) | 
| 412 | qed | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 413 | |
| 70136 | 414 | instance\<^marker>\<open>tag unimportant\<close> vec :: (perfect_space, finite) perfect_space | 
| 44571 | 415 | proof | 
| 416 |   fix x :: "'a ^ 'b" show "\<not> open {x}"
 | |
| 417 | proof | |
| 418 |     assume "open {x}"
 | |
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 419 |     hence "\<forall>i. open ((\<lambda>x. x $ i) ` {x})" by (fast intro: open_image_vec_nth)
 | 
| 44571 | 420 |     hence "\<forall>i. open {x $ i}" by simp
 | 
| 421 | thus "False" by (simp add: not_open_singleton) | |
| 422 | qed | |
| 423 | qed | |
| 424 | ||
| 425 | ||
| 69683 | 426 | subsection \<open>Metric space\<close> | 
| 62101 | 427 | (* 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: 
36590diff
changeset | 428 | |
| 70136 | 429 | 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: 
36590diff
changeset | 430 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 431 | |
| 70136 | 432 | definition\<^marker>\<open>tag important\<close> | 
| 67155 | 433 | "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: 
36590diff
changeset | 434 | |
| 62101 | 435 | instance .. | 
| 436 | end | |
| 437 | ||
| 70136 | 438 | instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) uniformity_dist | 
| 62101 | 439 | begin | 
| 440 | ||
| 70136 | 441 | definition\<^marker>\<open>tag important\<close> [code del]: | 
| 67731 | 442 |   "(uniformity :: (('a^'b::_) \<times> ('a^'b::_)) filter) =
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 443 |     (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
 | 
| 62101 | 444 | |
| 70136 | 445 | instance\<^marker>\<open>tag unimportant\<close> | 
| 62101 | 446 | by standard (rule uniformity_vec_def) | 
| 447 | end | |
| 448 | ||
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 449 | 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: 
62101diff
changeset | 450 | |
| 70136 | 451 | instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) metric_space | 
| 62101 | 452 | begin | 
| 453 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 454 | proposition dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y" | 
| 67155 | 455 | 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: 
36590diff
changeset | 456 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 457 | instance proof | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 458 | fix x y :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 459 | show "dist x y = 0 \<longleftrightarrow> x = y" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 460 | unfolding dist_vec_def | 
| 67155 | 461 | 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: 
36590diff
changeset | 462 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 463 | fix x y z :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 464 | 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: 
44135diff
changeset | 465 | unfolding dist_vec_def | 
| 67155 | 466 | apply (rule order_trans [OF _ L2_set_triangle_ineq]) | 
| 467 | apply (simp add: L2_set_mono dist_triangle2) | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 468 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 469 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 470 |   fix S :: "('a ^ 'b) set"
 | 
| 62101 | 471 | have *: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" | 
| 44630 | 472 | proof | 
| 473 | assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" | |
| 474 | proof | |
| 475 | fix x assume "x \<in> S" | |
| 476 | obtain A where A: "\<forall>i. open (A i)" "\<forall>i. x $ i \<in> A i" | |
| 477 | and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" | |
| 60420 | 478 | using \<open>open S\<close> and \<open>x \<in> S\<close> unfolding open_vec_def by metis | 
| 44630 | 479 | have "\<forall>i\<in>UNIV. \<exists>r>0. \<forall>y. dist y (x $ i) < r \<longrightarrow> y \<in> A i" | 
| 480 | using A unfolding open_dist by simp | |
| 481 | 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: 
44631diff
changeset | 482 | by (rule finite_set_choice [OF finite]) | 
| 44630 | 483 | then obtain r where r1: "\<forall>i. 0 < r i" | 
| 484 | and r2: "\<forall>i y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i" by fast | |
| 485 | have "0 < Min (range r) \<and> (\<forall>y. dist y x < Min (range r) \<longrightarrow> y \<in> S)" | |
| 486 | by (simp add: r1 r2 S le_less_trans [OF dist_vec_nth_le]) | |
| 487 | thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" .. | |
| 488 | qed | |
| 489 | next | |
| 490 | assume *: "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" show "open S" | |
| 491 | proof (unfold open_vec_def, rule) | |
| 492 | fix x assume "x \<in> S" | |
| 493 | then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S" | |
| 494 | using * by fast | |
| 63040 | 495 |       define r where [abs_def]: "r i = e / sqrt (of_nat CARD('b))" for i :: 'b
 | 
| 60420 | 496 | from \<open>0 < e\<close> have r: "\<forall>i. 0 < r i" | 
| 56541 | 497 | unfolding r_def by simp_all | 
| 67155 | 498 | from \<open>0 < e\<close> have e: "e = L2_set r UNIV" | 
| 499 | unfolding r_def by (simp add: L2_set_constant) | |
| 63040 | 500 |       define A where "A i = {y. dist (x $ i) y < r i}" for i
 | 
| 44630 | 501 | have "\<forall>i. open (A i) \<and> x $ i \<in> A i" | 
| 502 | unfolding A_def by (simp add: open_ball r) | |
| 503 | moreover have "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" | |
| 67155 | 504 | by (simp add: A_def S dist_vec_def e L2_set_strict_mono dist_commute) | 
| 44630 | 505 | ultimately show "\<exists>A. (\<forall>i. open (A i) \<and> x $ i \<in> A i) \<and> | 
| 506 | (\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S)" by metis | |
| 507 | qed | |
| 508 | qed | |
| 62101 | 509 | show "open S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)" | 
| 510 | unfolding * eventually_uniformity_metric | |
| 511 | 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: 
36590diff
changeset | 512 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 513 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 514 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 515 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 516 | lemma Cauchy_vec_nth: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 517 | "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: 
44135diff
changeset | 518 | 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: 
36590diff
changeset | 519 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 520 | lemma vec_CauchyI: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 521 | fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 522 | 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: 
36590diff
changeset | 523 | 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: 
69684diff
changeset | 524 | proof (rule metric_CauchyI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 525 | fix r :: real assume "0 < r" | 
| 56541 | 526 |   hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp
 | 
| 63040 | 527 | 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 | 
| 528 | define M where "M = Max (range N)" | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 529 | have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s" | 
| 60420 | 530 | 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: 
36590diff
changeset | 531 | 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: 
36590diff
changeset | 532 | unfolding N_def by (rule LeastI_ex) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 533 | 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: 
36590diff
changeset | 534 | unfolding M_def by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 535 |   {
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 536 | fix m n :: nat | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 537 | assume "M \<le> m" "M \<le> n" | 
| 67155 | 538 | 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: 
44135diff
changeset | 539 | unfolding dist_vec_def .. | 
| 64267 | 540 | also have "\<dots> \<le> sum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV" | 
| 67155 | 541 | by (rule L2_set_le_sum [OF zero_le_dist]) | 
| 64267 | 542 | also have "\<dots> < sum (\<lambda>i::'n. ?s) UNIV" | 
| 543 | 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: 
36590diff
changeset | 544 | also have "\<dots> = r" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 545 | by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 546 | finally have "dist (X m) (X n) < r" . | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 547 | } | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 548 | 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: 
36590diff
changeset | 549 | by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 550 | 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: 
36590diff
changeset | 551 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 552 | |
| 70136 | 553 | 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: 
36590diff
changeset | 554 | proof | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 555 | fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X" | 
| 61969 | 556 | have "\<And>i. (\<lambda>n. X n $ i) \<longlonglongrightarrow> lim (\<lambda>n. X n $ i)" | 
| 60420 | 557 | 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: 
36590diff
changeset | 558 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | 
| 61969 | 559 | 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: 
44135diff
changeset | 560 | by (simp add: vec_tendstoI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 561 | then show "convergent X" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 562 | by (rule convergentI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 563 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 564 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 565 | |
| 69683 | 566 | subsection \<open>Normed vector space\<close> | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 567 | |
| 70136 | 568 | 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: 
36590diff
changeset | 569 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 570 | |
| 70136 | 571 | 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: 
36590diff
changeset | 572 | |
| 70136 | 573 | 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: 
36590diff
changeset | 574 | |
| 70136 | 575 | instance\<^marker>\<open>tag unimportant\<close> proof | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 576 | fix a :: real and x y :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 577 | show "norm x = 0 \<longleftrightarrow> x = 0" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 578 | unfolding norm_vec_def | 
| 67155 | 579 | 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: 
36590diff
changeset | 580 | 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: 
44135diff
changeset | 581 | unfolding norm_vec_def | 
| 67155 | 582 | apply (rule order_trans [OF _ L2_set_triangle_ineq]) | 
| 583 | 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: 
36590diff
changeset | 584 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 585 | 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: 
44135diff
changeset | 586 | unfolding norm_vec_def | 
| 67155 | 587 | by (simp add: L2_set_right_distrib) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 588 | show "sgn x = scaleR (inverse (norm x)) x" | 
| 44141 | 589 | by (rule sgn_vec_def) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 590 | show "dist x y = norm (x - y)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 591 | unfolding dist_vec_def norm_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 592 | by (simp add: dist_norm) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 593 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 594 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 595 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 596 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 597 | 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: 
44135diff
changeset | 598 | unfolding norm_vec_def | 
| 67155 | 599 | by (rule member_le_L2_set) simp_all | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 600 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 601 | lemma norm_le_componentwise_cart: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 602 | 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: 
67982diff
changeset | 603 | 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: 
67982diff
changeset | 604 | shows "norm x \<le> norm y" | 
| 70136 | 605 | unfolding norm_vec_def | 
| 606 | 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: 
67982diff
changeset | 607 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 608 | lemma component_le_norm_cart: "\<bar>x$i\<bar> \<le> norm x" | 
| 73976 | 609 | 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: 
67982diff
changeset | 610 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 611 | 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: 
67982diff
changeset | 612 | 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: 
67982diff
changeset | 613 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 614 | 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: 
67982diff
changeset | 615 | 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: 
67982diff
changeset | 616 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 617 | 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: 
67982diff
changeset | 618 | 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: 
67982diff
changeset | 619 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 620 | lemma bounded_linear_vec_nth[intro]: "bounded_linear (\<lambda>x. x $ i)" | 
| 73976 | 621 | proof | 
| 622 | show "\<exists>K. \<forall>x. norm (x $ i) \<le> norm x * K" | |
| 623 | by (metis mult.commute mult.left_neutral norm_nth_le) | |
| 624 | qed auto | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 625 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 626 | instance vec :: (banach, finite) banach .. | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 627 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 628 | |
| 69683 | 629 | subsection \<open>Inner product space\<close> | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 630 | |
| 70136 | 631 | 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: 
36590diff
changeset | 632 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 633 | |
| 70136 | 634 | 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: 
36590diff
changeset | 635 | |
| 70136 | 636 | instance\<^marker>\<open>tag unimportant\<close> proof | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 637 | fix r :: real and x y z :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 638 | show "inner x y = inner y x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 639 | unfolding inner_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 640 | by (simp add: inner_commute) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 641 | 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: 
44135diff
changeset | 642 | unfolding inner_vec_def | 
| 64267 | 643 | by (simp add: inner_add_left sum.distrib) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 644 | 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: 
44135diff
changeset | 645 | unfolding inner_vec_def | 
| 64267 | 646 | by (simp add: sum_distrib_left) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 647 | show "0 \<le> inner x x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 648 | unfolding inner_vec_def | 
| 64267 | 649 | by (simp add: sum_nonneg) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 650 | show "inner x x = 0 \<longleftrightarrow> x = 0" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 651 | unfolding inner_vec_def | 
| 64267 | 652 | 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: 
36590diff
changeset | 653 | show "norm x = sqrt (inner x x)" | 
| 67155 | 654 | 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: 
36590diff
changeset | 655 | by (simp add: power2_norm_eq_inner) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 656 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 657 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 658 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 659 | |
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 660 | |
| 69683 | 661 | subsection \<open>Euclidean space\<close> | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 662 | |
| 60420 | 663 | text \<open>Vectors pointing along a single axis.\<close> | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 664 | |
| 70136 | 665 | 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: 
44165diff
changeset | 666 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 667 | lemma axis_nth [simp]: "axis i x $ i = x" | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 668 | unfolding axis_def by simp | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 669 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 670 | 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: 
44165diff
changeset | 671 | unfolding axis_def vec_eq_iff by auto | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 672 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 673 | lemma inner_axis_axis: | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 674 | "inner (axis i x) (axis j y) = (if i = j then inner x y else 0)" | 
| 73976 | 675 | 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: 
44165diff
changeset | 676 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 677 | 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: 
67968diff
changeset | 678 | 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: 
44165diff
changeset | 679 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 680 | 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: 
67673diff
changeset | 681 | by (simp add: inner_axis inner_commute) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 682 | |
| 70136 | 683 | 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: 
42290diff
changeset | 684 | begin | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 685 | |
| 70136 | 686 | 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: 
44165diff
changeset | 687 | |
| 70136 | 688 | instance\<^marker>\<open>tag unimportant\<close> proof | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 689 |   show "(Basis :: ('a ^ 'b) set) \<noteq> {}"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 690 | unfolding Basis_vec_def by simp | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 691 | next | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 692 |   show "finite (Basis :: ('a ^ 'b) set)"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 693 | unfolding Basis_vec_def by simp | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 694 | next | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 695 | fix u v :: "'a ^ 'b" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 696 | assume "u \<in> Basis" and "v \<in> Basis" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 697 | thus "inner u v = (if u = v then 1 else 0)" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 698 | unfolding Basis_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 699 | 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: 
42290diff
changeset | 700 | next | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 701 | fix x :: "'a ^ 'b" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 702 | show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 703 | unfolding Basis_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 704 | 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: 
50252diff
changeset | 705 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50252diff
changeset | 706 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 707 | 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: 
69684diff
changeset | 708 | proof - | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 709 |   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: 
67968diff
changeset | 710 | 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: 
67968diff
changeset | 711 |   also have "... = CARD('b) * DIM('a)"
 | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 712 | 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: 
67968diff
changeset | 713 | finally show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 714 | by (simp add: Basis_vec_def) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 715 | qed | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 716 | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 717 | end | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 718 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 719 | 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: 
67982diff
changeset | 720 | 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: 
67982diff
changeset | 721 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 722 | 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: 
67982diff
changeset | 723 | fixes f:: "'a \<Rightarrow> real ^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 724 | 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: 
67982diff
changeset | 725 |   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: 
69684diff
changeset | 726 | 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: 
69684diff
changeset | 727 | by simp | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 728 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 729 | 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: 
62102diff
changeset | 730 | by (simp add: inner_axis) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62102diff
changeset | 731 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 732 | 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: 
67968diff
changeset | 733 | 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: 
67968diff
changeset | 734 | 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: 
67968diff
changeset | 735 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 736 | 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: 
67968diff
changeset | 737 | 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: 
67968diff
changeset | 738 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 739 | 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: 
67968diff
changeset | 740 | 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: 
67968diff
changeset | 741 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 742 | lemma axis_inverse: | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 743 | fixes v :: "real^'n" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 744 | assumes "v \<in> Basis" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 745 | 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: 
67968diff
changeset | 746 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 747 |   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: 
67968diff
changeset | 748 | 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: 
67968diff
changeset | 749 | then show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 750 | 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: 
67968diff
changeset | 751 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 752 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 753 | lemma axis_index: | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 754 | fixes v :: "real^'n" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 755 | assumes "v \<in> Basis" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 756 | 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: 
67968diff
changeset | 757 | 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: 
67968diff
changeset | 758 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 759 | lemma axis_index_axis [simp]: | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 760 | fixes UU :: "real^'n" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67968diff
changeset | 761 | 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: 
67968diff
changeset | 762 | 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: 
62102diff
changeset | 763 | |
| 70136 | 764 | 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: 
67982diff
changeset | 765 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 766 | lemma sum_cong_aux: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 767 | "(\<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: 
67982diff
changeset | 768 | by (auto intro: sum.cong) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 769 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 770 | hide_fact (open) sum_cong_aux | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 771 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 772 | method_setup vector = \<open> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 773 | let | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 774 | val ss1 = | 
| 69597 | 775 | 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: 
67982diff
changeset | 776 |       addsimps [@{thm sum.distrib} RS sym,
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 777 |       @{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: 
67982diff
changeset | 778 |       @{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: 
67982diff
changeset | 779 | val ss2 = | 
| 69597 | 780 | simpset_of (\<^context> addsimps | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 781 |              [@{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: 
67982diff
changeset | 782 |               @{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: 
67982diff
changeset | 783 |               @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def},
 | 
| 71174 | 784 |               @{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: 
67982diff
changeset | 785 | fun vector_arith_tac ctxt ths = | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 786 | simp_tac (put_simpset ss1 ctxt) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 787 |     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: 
67982diff
changeset | 788 |          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: 
67982diff
changeset | 789 |          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: 
67982diff
changeset | 790 |     (* 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: 
67982diff
changeset | 791 | 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: 
67982diff
changeset | 792 | in | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 793 | 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: 
42290diff
changeset | 794 | end | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 795 | \<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: 
67982diff
changeset | 796 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 797 | 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: 
69684diff
changeset | 798 | 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: 
67982diff
changeset | 799 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 800 | 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: 
67982diff
changeset | 801 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 802 | 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: 
67982diff
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: 
69684diff
changeset | 804 | 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: 
69684diff
changeset | 805 | 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: 
69684diff
changeset | 806 | 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: 
69684diff
changeset | 807 | 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: 
67982diff
changeset | 808 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 809 | 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: 
67982diff
changeset | 810 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 811 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 812 | lemma vec_sum: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 813 | assumes "finite S" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 814 | 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: 
67982diff
changeset | 815 | using assms | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 816 | proof induct | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 817 | case empty | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 818 | then show ?case by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 819 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 820 | case insert | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 821 | 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: 
67982diff
changeset | 822 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 823 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 824 | text\<open>Obvious "component-pushing".\<close> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 825 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 826 | 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: 
67982diff
changeset | 827 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 828 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 829 | 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: 
67982diff
changeset | 830 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 831 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 832 | 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: 
67982diff
changeset | 833 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 834 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 835 | 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: 
67982diff
changeset | 836 | |
| 70136 | 837 | 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: 
67982diff
changeset | 838 | 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: 
67982diff
changeset | 839 | 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: 
67982diff
changeset | 840 | vector_scaleR_component cond_component | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 841 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 842 | |
| 70136 | 843 | 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: 
67982diff
changeset | 844 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 845 | 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: 
67982diff
changeset | 846 | by standard (vector mult.assoc) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 847 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 848 | 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: 
67982diff
changeset | 849 | by standard vector+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 850 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 851 | 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: 
67982diff
changeset | 852 | by standard (vector mult.commute) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 853 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 854 | 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: 
67982diff
changeset | 855 | by standard vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 856 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 857 | instance vec :: (semiring, finite) semiring | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 858 | by standard (vector field_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 859 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 860 | 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: 
67982diff
changeset | 861 | by standard (vector field_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 862 | 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: 
67982diff
changeset | 863 | by standard vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 864 | 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: 
67982diff
changeset | 865 | by standard (vector field_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 866 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 867 | 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: 
67982diff
changeset | 868 | 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: 
67982diff
changeset | 869 | 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: 
67982diff
changeset | 870 | instance vec :: (ring, finite) ring .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 871 | 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: 
67982diff
changeset | 872 | 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: 
67982diff
changeset | 873 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 874 | 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: 
67982diff
changeset | 875 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 876 | 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: 
67982diff
changeset | 877 | 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: 
67982diff
changeset | 878 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 879 | 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: 
67982diff
changeset | 880 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 881 | 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: 
67982diff
changeset | 882 | proof (induct n) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 883 | case 0 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 884 | then show ?case by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 885 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 886 | case Suc | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 887 | then show ?case by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 888 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 889 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 890 | 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: 
67982diff
changeset | 891 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 892 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 893 | 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: 
67982diff
changeset | 894 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 895 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 896 | 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: 
67982diff
changeset | 897 | proof | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 898 | fix m n :: nat | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 899 | 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: 
67982diff
changeset | 900 | 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: 
67982diff
changeset | 901 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 902 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 903 | instance vec :: (numeral, finite) numeral .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 904 | 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: 
67982diff
changeset | 905 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 906 | 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: 
67982diff
changeset | 907 | 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: 
67982diff
changeset | 908 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 909 | 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: 
67982diff
changeset | 910 | 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: 
67982diff
changeset | 911 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 912 | 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: 
67982diff
changeset | 913 | 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: 
67982diff
changeset | 914 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 915 | 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: 
67982diff
changeset | 916 | by (vector mult.assoc) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 917 | 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: 
67982diff
changeset | 918 | by (vector field_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 919 | 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: 
67982diff
changeset | 920 | by (vector field_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 921 | 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: 
67982diff
changeset | 922 | 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: 
67982diff
changeset | 923 | 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: 
67982diff
changeset | 924 | by (vector field_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 925 | 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: 
67982diff
changeset | 926 | 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: 
67982diff
changeset | 927 | 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: 
67982diff
changeset | 928 | 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: 
67982diff
changeset | 929 | 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: 
67982diff
changeset | 930 | by (vector field_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 931 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 932 | 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: 
67982diff
changeset | 933 | by (simp add: vec_eq_iff) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 934 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 935 | 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: 
67982diff
changeset | 936 | by unfold_locales (vector algebra_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 937 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 938 | 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: 
67982diff
changeset | 939 | by vector | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 940 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 941 | 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: 
67982diff
changeset | 942 | 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: 
67982diff
changeset | 943 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 944 | 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: 
71174diff
changeset | 945 | 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: 
67982diff
changeset | 946 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 947 | 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: 
67982diff
changeset | 948 | 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: 
67982diff
changeset | 949 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 950 | 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: 
67982diff
changeset | 951 | 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: 
67982diff
changeset | 952 | 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: 
67982diff
changeset | 953 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 954 | lemma sum_component [simp]: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 955 |   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: 
67982diff
changeset | 956 | 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: 
67982diff
changeset | 957 | proof (cases "finite S") | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 958 | case True | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 959 | 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: 
67982diff
changeset | 960 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 961 | case False | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 962 | then show ?thesis by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 963 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 964 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 965 | 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: 
67982diff
changeset | 966 | by (simp add: vec_eq_iff) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 967 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 968 | lemma sum_cmul: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 969 |   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: 
67982diff
changeset | 970 | 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: 
67982diff
changeset | 971 | 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: 
67982diff
changeset | 972 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 973 | lemma linear_vec [simp]: "linear vec" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 974 | using Vector_Spaces_linear_vec | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 975 | by unfold_locales (vector algebra_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 976 | |
| 69683 | 977 | subsection \<open>Matrix operations\<close> | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 978 | |
| 69597 | 979 | 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: 
67982diff
changeset | 980 | |
| 70136 | 981 | 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: 
67982diff
changeset | 982 | "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: 
67982diff
changeset | 983 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 984 | 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: 
67982diff
changeset | 985 | by (simp add: map_matrix_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 986 | |
| 70136 | 987 | definition\<^marker>\<open>tag important\<close> matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 988 | (infixl "**" 70) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 989 | 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: 
67982diff
changeset | 990 | |
| 70136 | 991 | definition\<^marker>\<open>tag important\<close> matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 992 | (infixl "*v" 70) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 993 | 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: 
67982diff
changeset | 994 | |
| 70136 | 995 | definition\<^marker>\<open>tag important\<close> vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 996 | (infixl "v*" 70) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 997 | where "v v* m == (\<chi> j. sum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 998 | |
| 70136 | 999 | definition\<^marker>\<open>tag unimportant\<close> "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)" | 
| 1000 | 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: 
67982diff
changeset | 1001 | "(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))" | 
| 70136 | 1002 | definition\<^marker>\<open>tag unimportant\<close> "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))" | 
| 1003 | definition\<^marker>\<open>tag unimportant\<close> "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))" | |
| 1004 | definition\<^marker>\<open>tag unimportant\<close> "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
 | |
| 1005 | 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: 
67982diff
changeset | 1006 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1007 | 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: 
68072diff
changeset | 1008 | 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: 
68072diff
changeset | 1009 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1010 | 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: 
68072diff
changeset | 1011 | 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: 
68072diff
changeset | 1012 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1013 | 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: 
69684diff
changeset | 1014 | 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: 
67982diff
changeset | 1015 | 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: 
67982diff
changeset | 1016 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1017 | lemma matrix_mul_lid [simp]: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1018 | 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: 
67982diff
changeset | 1019 | shows "mat 1 ** A = A" | 
| 73976 | 1020 | unfolding matrix_matrix_mult_def mat_def | 
| 1021 | 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: 
67982diff
changeset | 1022 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1023 | lemma matrix_mul_rid [simp]: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1024 | 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: 
67982diff
changeset | 1025 | shows "A ** mat 1 = A" | 
| 73976 | 1026 | unfolding matrix_matrix_mult_def mat_def | 
| 1027 | 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: 
67982diff
changeset | 1028 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1029 | proposition 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: 
67982diff
changeset | 1030 | apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1031 | apply (subst sum.swap) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1032 | apply simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1033 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1034 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1035 | proposition 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: 
67982diff
changeset | 1036 | 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: 
67982diff
changeset | 1037 | sum_distrib_left sum_distrib_right mult.assoc) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1038 | apply (subst sum.swap) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1039 | apply simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1040 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1041 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1042 | proposition scalar_matrix_assoc: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1043 |   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: 
68072diff
changeset | 1044 | 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: 
68072diff
changeset | 1045 | 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: 
68072diff
changeset | 1046 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1047 | proposition matrix_scalar_ac: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1048 |   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: 
68072diff
changeset | 1049 | 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: 
68072diff
changeset | 1050 | 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: 
68072diff
changeset | 1051 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1052 | 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: 
67982diff
changeset | 1053 | apply (vector matrix_vector_mult_def mat_def) | 
| 71044 | 1054 | 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: 
67982diff
changeset | 1055 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
changeset | 1061 | lemma matrix_mult_transpose_dot_column: | 
| 69669 | 1062 | shows "transpose A ** A = (\<chi> i j. inner (column i A) (column j A))" | 
| 1063 | by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def) | |
| 1064 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1065 | lemma matrix_mult_transpose_dot_row: | 
| 69669 | 1066 | shows "A ** transpose A = (\<chi> i j. inner (row i A) (row j A))" | 
| 1067 | by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def) | |
| 1068 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1069 | lemma matrix_eq: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1070 | fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" | 
| 73976 | 1071 | shows "A = B \<longleftrightarrow> (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs") | 
| 1072 | proof | |
| 1073 | assume ?rhs | |
| 1074 | then show ?lhs | |
| 1075 | apply (subst vec_eq_iff) | |
| 1076 | apply (clarsimp simp add: matrix_vector_mult_def if_distrib if_distribR vec_eq_iff cong: if_cong) | |
| 1077 | apply (erule_tac x="axis ia 1" in allE) | |
| 1078 | apply (erule_tac x="i" in allE) | |
| 1079 | apply (auto simp add: if_distrib if_distribR axis_def | |
| 1080 | sum.delta[OF finite] cong del: if_weak_cong) | |
| 1081 | done | |
| 1082 | qed auto | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
changeset | 1089 | apply (subst sum.swap) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1090 | apply simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1091 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
changeset | 1097 | by (vector transpose_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
68072diff
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: 
68072diff
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: 
68072diff
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: 
69684diff
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: 
68072diff
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: 
68072diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
changeset | 1122 | lemma vector_componentwise: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
changeset | 1128 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1129 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
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: 
67982diff
changeset | 1131 | |
| 70136 | 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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
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: 
67982diff
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: 
67982diff
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: 
69684diff
changeset | 1141 | lemma matrix_vector_mul_linear[intro, simp]: "linear (\<lambda>x. A *v (x::'a::real_algebra_1 ^ _))" | 
| 68074 | 1142 | by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff field_simps sum_distrib_left | 
| 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: 
67982diff
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: 
69684diff
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: 
68072diff
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: 
68072diff
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: 
68072diff
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: 
68072diff
changeset | 1149 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1150 | 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: 
67982diff
changeset | 1151 | "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: 
67982diff
changeset | 1152 | 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: 
67982diff
changeset | 1153 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1154 | 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: 
67982diff
changeset | 1155 | 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: 
67982diff
changeset | 1156 | 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: 
67982diff
changeset | 1157 | 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: 
67982diff
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: 
69684diff
changeset | 1159 | 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: 
67982diff
changeset | 1160 | fixes A :: "real^'n^'m" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1161 | 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: 
67982diff
changeset | 1162 | 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: 
67982diff
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: 
69684diff
changeset | 1164 | 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: 
67982diff
changeset | 1165 | 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: 
67982diff
changeset | 1166 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1167 | 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: 
67982diff
changeset | 1168 | 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: 
67982diff
changeset | 1169 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1170 | 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: 
67982diff
changeset | 1171 | "(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: 
67982diff
changeset | 1172 | 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: 
67982diff
changeset | 1173 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1174 | 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: 
67982diff
changeset | 1175 | 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: 
67982diff
changeset | 1176 | 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: 
67982diff
changeset | 1177 | 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: 
67982diff
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: 
69684diff
changeset | 1179 | lemma matrix_vector_column: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1180 | "(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: 
67982diff
changeset | 1181 | 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: 
67982diff
changeset | 1182 | |
| 69683 | 1183 | 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: 
67982diff
changeset | 1184 | |
| 70136 | 1185 | 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: 
67982diff
changeset | 1186 | "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: 
67982diff
changeset | 1187 | |
| 70136 | 1188 | 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: 
67982diff
changeset | 1189 | "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: 
67982diff
changeset | 1190 | (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: 
67982diff
changeset | 1191 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1192 | lemma inj_matrix_vector_mult: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1193 | fixes A::"'a::field^'n^'m" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1194 | assumes "invertible A" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1195 | shows "inj ((*v) A)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1196 | 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: 
67982diff
changeset | 1197 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1198 | lemma scalar_invertible: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1199 |   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: 
68072diff
changeset | 1200 | 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: 
68072diff
changeset | 1201 | 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: 
69684diff
changeset | 1202 | proof - | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1203 | 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: 
68072diff
changeset | 1204 | using assms unfolding invertible_def by auto | 
| 69272 | 1205 | 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: 
68072diff
changeset | 1206 | 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: 
68072diff
changeset | 1207 | 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: 
68072diff
changeset | 1208 | 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: 
68072diff
changeset | 1209 | 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: 
68072diff
changeset | 1210 | qed | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
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: 
69684diff
changeset | 1212 | proposition scalar_invertible_iff: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
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: 
68072diff
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: 
68072diff
changeset | 1215 | 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: 
68072diff
changeset | 1216 | 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: 
68072diff
changeset | 1217 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1218 | lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x" | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1219 | unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1220 | by simp | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1221 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1222 | lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A" | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1223 | unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1224 | by simp | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1225 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1226 | 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: 
68072diff
changeset | 1227 |   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: 
68072diff
changeset | 1228 | 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: 
68072diff
changeset | 1229 | 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: 
68072diff
changeset | 1230 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1231 | 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: 
68072diff
changeset | 1232 |   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: 
68072diff
changeset | 1233 | 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: 
68072diff
changeset | 1234 | 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: 
68072diff
changeset | 1235 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1236 | 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: 
68072diff
changeset | 1237 | 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: 
68072diff
changeset | 1238 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1239 | 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: 
68072diff
changeset | 1240 | 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: 
68072diff
changeset | 1241 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1242 | lemma vector_matrix_mul_rid [simp]: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1243 |   fixes v :: "('a::semiring_1)^'n"
 | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1244 | shows "v v* mat 1 = v" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1245 | by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1246 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1247 | 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: 
68072diff
changeset | 1248 | 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: 
68072diff
changeset | 1249 | 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: 
68072diff
changeset | 1250 | 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: 
68072diff
changeset | 1251 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69684diff
changeset | 1252 | proposition 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: 
68072diff
changeset | 1253 | 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: 
68072diff
changeset | 1254 | 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: 
69684diff
changeset | 1255 | proof - | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1256 | have "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: 
68072diff
changeset | 1257 | 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: 
68072diff
changeset | 1258 | by (simp add: algebra_simps) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1259 | 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: 
68072diff
changeset | 1260 | 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: 
68072diff
changeset | 1261 | by auto | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1262 | qed | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 1263 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1264 | end |