author | wenzelm |
Sat, 11 Nov 2023 20:01:14 +0100 | |
changeset 78943 | bc89bdc65f29 |
parent 70136 | f03a01a18c6e |
permissions | -rw-r--r-- |
63627 | 1 |
(* Title: HOL/Analysis/Norm_Arith.thy |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
2 |
Author: Amine Chaieb, University of Cambridge |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
3 |
*) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
4 |
|
69517 | 5 |
section \<open>Linear Decision Procedure for Normed Spaces\<close> |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
6 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
7 |
theory Norm_Arith |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63928
diff
changeset
|
8 |
imports "HOL-Library.Sum_of_Squares" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
9 |
begin |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
10 |
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11 |
(* FIXME: move elsewhere *) |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
12 |
lemma sum_sqs_eq: |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13 |
fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \<Longrightarrow> y = x" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
14 |
by algebra |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
15 |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
16 |
lemma norm_cmul_rule_thm: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
17 |
fixes x :: "'a::real_normed_vector" |
53717 | 18 |
shows "b \<ge> norm x \<Longrightarrow> \<bar>c\<bar> * b \<ge> norm (scaleR c x)" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
19 |
unfolding norm_scaleR |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
20 |
apply (erule mult_left_mono) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
21 |
apply simp |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
22 |
done |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
23 |
|
53717 | 24 |
(* FIXME: Move all these theorems into the ML code using lemma antiquotation *) |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
25 |
lemma norm_add_rule_thm: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
26 |
fixes x1 x2 :: "'a::real_normed_vector" |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
27 |
shows "norm x1 \<le> b1 \<Longrightarrow> norm x2 \<le> b2 \<Longrightarrow> norm (x1 + x2) \<le> b1 + b2" |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
28 |
by (rule order_trans [OF norm_triangle_ineq add_mono]) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
29 |
|
53717 | 30 |
lemma ge_iff_diff_ge_0: |
31 |
fixes a :: "'a::linordered_ring" |
|
32 |
shows "a \<ge> b \<equiv> a - b \<ge> 0" |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
33 |
by (simp add: field_simps) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
34 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
35 |
lemma pth_1: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
36 |
fixes x :: "'a::real_normed_vector" |
53717 | 37 |
shows "x \<equiv> scaleR 1 x" by simp |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
38 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
39 |
lemma pth_2: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
40 |
fixes x :: "'a::real_normed_vector" |
53717 | 41 |
shows "x - y \<equiv> x + -y" |
42 |
by (atomize (full)) simp |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
43 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
44 |
lemma pth_3: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
45 |
fixes x :: "'a::real_normed_vector" |
53717 | 46 |
shows "- x \<equiv> scaleR (-1) x" |
47 |
by simp |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
48 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
49 |
lemma pth_4: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
50 |
fixes x :: "'a::real_normed_vector" |
53717 | 51 |
shows "scaleR 0 x \<equiv> 0" |
52 |
and "scaleR c 0 = (0::'a)" |
|
53 |
by simp_all |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
54 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
55 |
lemma pth_5: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
56 |
fixes x :: "'a::real_normed_vector" |
53717 | 57 |
shows "scaleR c (scaleR d x) \<equiv> scaleR (c * d) x" |
58 |
by simp |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
59 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
60 |
lemma pth_6: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
61 |
fixes x :: "'a::real_normed_vector" |
53717 | 62 |
shows "scaleR c (x + y) \<equiv> scaleR c x + scaleR c y" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
63 |
by (simp add: scaleR_right_distrib) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
64 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
65 |
lemma pth_7: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
66 |
fixes x :: "'a::real_normed_vector" |
53717 | 67 |
shows "0 + x \<equiv> x" |
68 |
and "x + 0 \<equiv> x" |
|
69 |
by simp_all |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
70 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
71 |
lemma pth_8: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
72 |
fixes x :: "'a::real_normed_vector" |
53717 | 73 |
shows "scaleR c x + scaleR d x \<equiv> scaleR (c + d) x" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
74 |
by (simp add: scaleR_left_distrib) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
75 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
76 |
lemma pth_9: |
53717 | 77 |
fixes x :: "'a::real_normed_vector" |
78 |
shows "(scaleR c x + z) + scaleR d x \<equiv> scaleR (c + d) x + z" |
|
79 |
and "scaleR c x + (scaleR d x + z) \<equiv> scaleR (c + d) x + z" |
|
80 |
and "(scaleR c x + w) + (scaleR d x + z) \<equiv> scaleR (c + d) x + (w + z)" |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
81 |
by (simp_all add: algebra_simps) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
82 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
83 |
lemma pth_a: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
84 |
fixes x :: "'a::real_normed_vector" |
53717 | 85 |
shows "scaleR 0 x + y \<equiv> y" |
86 |
by simp |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
87 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
88 |
lemma pth_b: |
53717 | 89 |
fixes x :: "'a::real_normed_vector" |
90 |
shows "scaleR c x + scaleR d y \<equiv> scaleR c x + scaleR d y" |
|
91 |
and "(scaleR c x + z) + scaleR d y \<equiv> scaleR c x + (z + scaleR d y)" |
|
92 |
and "scaleR c x + (scaleR d y + z) \<equiv> scaleR c x + (scaleR d y + z)" |
|
93 |
and "(scaleR c x + w) + (scaleR d y + z) \<equiv> scaleR c x + (w + (scaleR d y + z))" |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
94 |
by (simp_all add: algebra_simps) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
95 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
96 |
lemma pth_c: |
53717 | 97 |
fixes x :: "'a::real_normed_vector" |
98 |
shows "scaleR c x + scaleR d y \<equiv> scaleR d y + scaleR c x" |
|
99 |
and "(scaleR c x + z) + scaleR d y \<equiv> scaleR d y + (scaleR c x + z)" |
|
100 |
and "scaleR c x + (scaleR d y + z) \<equiv> scaleR d y + (scaleR c x + z)" |
|
101 |
and "(scaleR c x + w) + (scaleR d y + z) \<equiv> scaleR d y + ((scaleR c x + w) + z)" |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
102 |
by (simp_all add: algebra_simps) |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
103 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
104 |
lemma pth_d: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
105 |
fixes x :: "'a::real_normed_vector" |
53717 | 106 |
shows "x + 0 \<equiv> x" |
107 |
by simp |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
108 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
109 |
lemma norm_imp_pos_and_ge: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
110 |
fixes x :: "'a::real_normed_vector" |
53717 | 111 |
shows "norm x \<equiv> n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
112 |
by atomize auto |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
113 |
|
53717 | 114 |
lemma real_eq_0_iff_le_ge_0: |
115 |
fixes x :: real |
|
116 |
shows "x = 0 \<equiv> x \<ge> 0 \<and> - x \<ge> 0" |
|
117 |
by arith |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
118 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
119 |
lemma norm_pths: |
53717 | 120 |
fixes x :: "'a::real_normed_vector" |
121 |
shows "x = y \<longleftrightarrow> norm (x - y) \<le> 0" |
|
122 |
and "x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)" |
|
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
123 |
using norm_ge_zero[of "x - y"] by auto |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
124 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
125 |
lemmas arithmetic_simps = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
126 |
arith_simps |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
127 |
add_numeral_special |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
128 |
add_neg_numeral_special |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
129 |
mult_1_left |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
130 |
mult_1_right |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
44516
diff
changeset
|
131 |
|
69605 | 132 |
ML_file \<open>normarith.ML\<close> |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
133 |
|
70136 | 134 |
method_setup\<^marker>\<open>tag important\<close> norm = \<open> |
53717 | 135 |
Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac) |
60420 | 136 |
\<close> "prove simple linear statements about vector norms" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
137 |
|
53717 | 138 |
|
60420 | 139 |
text \<open>Hence more metric properties.\<close> |
70136 | 140 |
text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close> |
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
67962
diff
changeset
|
141 |
proposition dist_triangle_add: |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
142 |
fixes x y x' y' :: "'a::real_normed_vector" |
53717 | 143 |
shows "dist (x + y) (x' + y') \<le> dist x x' + dist y y'" |
44516
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
144 |
by norm |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
145 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
146 |
lemma dist_triangle_add_half: |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
147 |
fixes x x' y y' :: "'a::real_normed_vector" |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
148 |
shows "dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 \<Longrightarrow> dist(x + y) (x' + y') < e" |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
149 |
by norm |
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
150 |
|
d9a496ae5d9d
move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents:
diff
changeset
|
151 |
end |