-
Notifications
You must be signed in to change notification settings - Fork 0
/
diag_invertibility.v
321 lines (293 loc) · 8.31 KB
/
diag_invertibility.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
Require Import vcfloat.VCFloat vcfloat.FPStdLib.
From mathcomp Require Import all_ssreflect ssralg ssrnat all_algebra seq matrix .
From mathcomp.analysis Require Import Rstruct.
Require Import fma_real_func_model fma_floating_point_model.
Require Import inf_norm_properties.
Require Import Coq.Lists.List. Import ListNotations.
Set Bullet Behavior "Strict Subproofs".
Require Import lemmas.
Require Import jacobi_preconditions.
Section WITH_NANS.
Context {NANS: Nans}.
Lemma transpose_implies {n}:
forall v1 v2: 'rV[R]_n,
v1^T = v2^T ->
v1 = v2.
Proof.
intros.
apply matrixP. unfold eqrel. intros.
apply matrixP in H. unfold eqrel in H.
specialize (H y x).
by rewrite !mxE in H.
Qed.
Lemma transpose_implies_col {n}:
forall v1 v2: 'rV[R]_n.+1,
v1 = v2 ->
v1^T = v2^T.
Proof.
intros.
apply matrixP. unfold eqrel. intros.
apply matrixP in H. unfold eqrel in H.
specialize (H y x). by rewrite !mxE.
Qed.
Lemma transpose_idempotent {m n:nat} (A: 'M[R]_(m,n)):
A = (A^T)^T.
Proof.
apply matrixP. unfold eqrel. intros.
by rewrite !mxE.
Qed.
Lemma vec_norm_implies_vec_zero {n: nat}:
forall v: 'cV[R]_n.+1,
vec_inf_norm v = 0%Re ->
(forall i, Rabs (v i ord0) = 0%Re).
Proof.
intros.
apply Rge_ge_eq in H.
apply Rge_ge_eq.
destruct H as [Ha Hb].
split.
+ apply Rle_ge. apply Rabs_pos.
+ apply Rle_ge. apply Rge_le in Hb.
unfold vec_inf_norm in Hb.
pose proof bigmax_le_implies.
specialize (H 0%Re [seq Rabs (v i 0)
| i <- enum 'I_n.+1] 0%Re).
rewrite size_map size_enum_ord in H.
assert ((0 < n.+1)%nat). by [].
specialize (H H0 Hb i).
assert ((i < n.+1)%nat). by apply ltn_ord.
specialize (H H1). rewrite seq_equiv in H.
rewrite nth_mkseq in H. rewrite inord_val in H.
apply H. apply ltn_ord.
Qed.
Local Open Scope R_scope.
Lemma max_order:
forall x y:R,
Rmax x y = x \/ Rmax x y = y.
Proof.
intros.
assert ((x <= y)%Re \/ (y <= x)%Re). nra.
destruct H.
+ rewrite Rmax_right. by right. apply H.
+ rewrite Rmax_left. by left. apply H.
Qed.
Lemma bigmax_destruct (a x0:R) s:
bigmaxr x0 (a :: s) = a \/
bigmaxr x0 (a :: s) = bigmaxr x0 s.
Proof.
assert (s = [::] \/ s != [::]).
{ destruct s.
+ by left.
+ by right.
} destruct H.
+ rewrite H. rewrite bigmaxr_un. by left.
+ assert (s = seq.head x0 s :: behead s).
{ by apply s_destruct. } rewrite H0.
rewrite bigmaxr_cons. rewrite -H0.
rewrite -RmaxE. apply max_order.
Qed.
Lemma bigmax_not_0_implies_aux (x0:R) s:
(0 < size s)%nat ->
(exists i, (i < size s)%nat /\
seq.nth x0 s i = bigmaxr x0 s).
Proof.
intros.
induction s.
+ by simpl in H.
+ pose proof (bigmax_destruct a x0 s).
destruct H0.
- exists 0%nat. simpl. split.
by []. by rewrite H0.
- assert (s = [::] \/ s != [::]).
{ destruct s.
+ by left.
+ by right.
} destruct H1.
* rewrite H1. exists 0%nat.
simpl. split. by []. by rewrite bigmaxr_un.
* assert ((0 < size s)%nat).
{ destruct s. by []. by simpl. } specialize (IHs H2).
destruct IHs as [i [Hsize IHs]].
exists i.+1. split.
++ simpl. by [].
++ simpl. by rewrite H0.
Qed.
Lemma bigmax_not_0_implies (x0:R) s:
bigmaxr x0 s <> x0 ->
(exists i, (i < size s)%nat /\
seq.nth x0 s i = bigmaxr x0 s /\
seq.nth x0 s i <> x0).
Proof.
intros.
assert (s = [::] \/ s != [::]).
{ destruct s.
+ by left.
+ by right.
} destruct H0.
+ rewrite H0 in H. rewrite bigmaxr_nil in H.
by contradict H.
+ assert ((0 < size s)%nat).
{ destruct s. by []. by simpl. }
pose proof (bigmax_not_0_implies_aux x0 s H1).
destruct H2 as [i [Hsize H2]].
exists i. split.
apply Hsize. split.
apply H2. by rewrite H2.
Qed.
Close Scope R_scope.
Lemma R0_no_Rabs:
forall x:R,
Rabs x <> 0%Re -> x <> 0%Re.
Proof.
intros.
pose proof (Rcase_abs x). destruct H0.
+ nra.
+ assert (x = 0%Re \/ (0 < x)%Re). nra.
destruct H0.
- rewrite H0 in H. rewrite Rabs_R0 in H. nra.
- nra.
Qed.
Lemma vec_norm_not_zero_implies {n:nat}:
forall v: 'cV[R]_n.+1,
vec_inf_norm v <> 0%Re ->
exists k, Rabs (v k ord0) = vec_inf_norm v /\
v k ord0 <> 0%Re.
Proof.
intros. unfold vec_inf_norm in H.
pose proof bigmax_not_0_implies.
specialize (H0 0%Re [seq Rabs (v i ord0)
| i <- enum 'I_n.+1] H).
rewrite size_map size_enum_ord in H0.
destruct H0 as [i [Hsize H0]].
exists (@inord n i).
unfold vec_inf_norm.
rewrite seq_equiv in H0. rewrite nth_mkseq in H0; last by apply Hsize.
destruct H0 as [H0a H0b].
split.
+ rewrite seq_equiv. apply H0a.
+ by apply R0_no_Rabs.
Qed.
Lemma Rabs_sub:
forall x y:R,
(Rabs x - Rabs y <= Rabs (x + y))%Re.
Proof.
intros.
assert ((x + y)%Re = (x - (-y))%Re). nra.
rewrite H.
rewrite -[in X in (_ - X <= _)%Re]Rabs_Ropp.
apply Rabs_triang_inv.
Qed.
Lemma Rabs_ineq_filter {n:nat} (f : 'I_n.+1 -> R) P:
(Rabs (\sum_(j < n.+1 | P j) f j) <= \sum_(j < n.+1 | P j) (Rabs (f j)))%Re.
Proof.
intros. rewrite big_mkcond /=.
rewrite [in X in (_ <= X)%Re]big_mkcond /=.
induction n.
+ simpl. rewrite !big_ord_recr //= !big_ord0.
rewrite -!RplusE. rewrite !Rplus_0_l.
case: (P ord_max).
- apply Rle_refl.
- rewrite Rabs_R0. apply Rle_refl.
+ simpl. rewrite big_ord_recr //=.
rewrite [in X in (_ <= X)%Re]big_ord_recr //=.
rewrite -!RplusE.
eapply Rle_trans. apply Rabs_triang.
apply Rplus_le_compat.
- apply IHn.
- case: (P ord_max).
* apply Rle_refl.
* rewrite Rabs_R0. apply Rle_refl.
Qed.
Lemma pos_to_gt:
forall x:R, x <> 0%Re -> (0 <= x)%Re -> (0 < x)%Re.
Proof.
intros.
assert (x = 0%Re \/ (0 < x)%Re). nra.
destruct H1.
+ rewrite H1 in H. nra.
+ nra.
Qed.
Lemma Rabs_0_implies_0:
forall x:R,
Rabs x = 0%Re -> x = 0%Re.
Proof.
intros.
pose proof Rcase_abs. specialize (H0 x).
destruct H0.
+ rewrite Rabs_left in H; nra.
+ rewrite Rabs_right in H; nra.
Qed.
Lemma diagonal_dominance_implies_invertibility {t} {n:nat}
(A: 'M[ftype t]_n.+1):
strict_diagonal_dominance A ->
(FT2R_mat A) \in unitmx.
Proof.
intros.
rewrite -unitmx_tr.
unfold strict_diagonal_dominance in H.
rewrite -row_free_unit.
apply inj_row_free. intros.
pose proof (@transpose_implies_col n ).
specialize (H1 (v *m (FT2R_mat A)^T) 0 H0).
rewrite trmx_mul in H1.
rewrite -transpose_idempotent in H1.
apply transpose_implies.
remember (v^T) as v_c.
rewrite -Heqv_c. rewrite trmx0.
rewrite -Heqv_c trmx0 in H1.
assert (vec_inf_norm v_c = 0%Re \/ vec_inf_norm v_c <> 0%Re).
{ nra. } destruct H2.
+ apply matrixP. unfold eqrel. intros.
pose proof (@vec_norm_implies_vec_zero n v_c H2 x).
rewrite !mxE.
assert (y = ord0). by apply ord1. rewrite H4.
by apply Rabs_0_implies_0.
+ pose proof (@vec_norm_not_zero_implies n v_c H2).
destruct H3 as [k [Habs Hneq0]].
contradict H1.
apply /eqP. apply /cV0Pn.
exists k. rewrite !mxE. apply /eqP.
apply R0_no_Rabs.
assert (forall x:R, (0 < x)%Re -> x <> 0%Re).
{ intros. nra. } apply H1.
rewrite (bigD1 k) /=.
eapply Rlt_le_trans; last by apply Rabs_sub.
rewrite -RmultE. rewrite Rabs_mult.
eapply Rlt_le_trans; last by
(apply Rplus_le_compat_l;
apply Ropp_le_contravar; apply Rabs_ineq_filter).
assert (\sum_(j < n.+1 | (fun j0 : 'I_n.+1 =>
j0 != k) j)
Rabs ((fun j0 : 'I_n.+1 =>
FT2R_mat A k j0 * v_c j0 0) j) =
\sum_(j < n.+1 | j != k) (Rabs (FT2R_mat A k j) * Rabs (v_c j ord0))%Re).
{ apply eq_big. by []. intros. rewrite -RmultE.
by rewrite Rabs_mult.
} rewrite H3. rewrite Habs.
apply Rlt_le_trans with
(Rabs (FT2R_mat A k k) * vec_inf_norm v_c -
(\sum_(j < n.+1 | j != k) (Rabs (FT2R_mat A k j) *
vec_inf_norm v_c)%Re))%Re.
- rewrite -big_distrl /=. rewrite -RmultE.
rewrite -Rmult_minus_distr_r.
apply Rmult_lt_0_compat.
* specialize (H k). apply Rlt_Rminus.
apply Rgt_lt. apply H.
* apply pos_to_gt. apply H2.
apply /RleP. apply vec_norm_pd.
- apply Rplus_le_compat_l. apply Ropp_le_contravar.
apply /RleP. apply big_sum_ge_ex_abstract.
intros. apply Rmult_le_compat_l. apply Rabs_pos.
unfold vec_inf_norm.
apply Rle_trans with
[seq Rabs (v_c i0 0)
| i0 <- enum 'I_n.+1]`_i.
* rewrite seq_equiv. rewrite nth_mkseq; last by apply ltn_ord.
rewrite inord_val. apply Rle_refl.
* apply /RleP.
apply (@bigmaxr_ler _ 0%Re [seq Rabs (v_c i0 0)
| i0 <- enum 'I_n.+1] i).
rewrite size_map size_enum_ord. apply ltn_ord.
- auto.
Qed.
End WITH_NANS.