Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
N
Numerical Analysis in Coq
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Wiki
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
Snippets
Build
Pipelines
Jobs
Pipeline schedules
Artifacts
Deploy
Releases
Package Registry
Model registry
Operate
Environments
Terraform modules
Monitor
Incidents
Analyze
Value stream analytics
Contributor analytics
CI/CD analytics
Repository analytics
Model experiments
Help
Help
Support
GitLab documentation
Compare GitLab plans
Community forum
Contribute to GitLab
Provide feedback
Keyboard shortcuts
?
Snippets
Groups
Projects
Show more breadcrumbs
Micaela Mayero
Numerical Analysis in Coq
Commits
946336ec
Commit
946336ec
authored
1 year ago
by
François Clément
Browse files
Options
Downloads
Patches
Plain Diff
Alternate def for dual_basis.
Rename bidual_nat_isom* -> bidual_isom*.
parent
dddfd38c
No related branches found
No related tags found
No related merge requests found
Pipeline
#7234
waiting for manual action
Stage: test
Changes
1
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
FEM/Algebra/Finite_dim_R.v
+16
-7
16 additions, 7 deletions
FEM/Algebra/Finite_dim_R.v
with
16 additions
and
7 deletions
FEM/Algebra/Finite_dim_R.v
+
16
−
7
View file @
946336ec
...
@@ -1494,6 +1494,15 @@ Definition dual_basis : '(PE_ms -> R)^n :=
...
@@ -1494,6 +1494,15 @@ Definition dual_basis : '(PE_ms -> R)^n :=
fun
i
x
=>
fun
i
x
=>
let
Hx
:=
eq_ind_r
(
@^~
(
val
x
))
(
in_sub
x
)
(
eq_sym
(
proj1
HB
))
in
let
Hx
:=
eq_ind_r
(
@^~
(
val
x
))
(
in_sub
x
)
(
eq_sym
(
proj1
HB
))
in
proj1_sig
(
span_EX
_
_
Hx
)
i
.
proj1_sig
(
span_EX
_
_
Hx
)
i
.
(
*
Proof
.
intros
i
[
x
Hx
].
rewrite
(
proj1
HB
)
in
Hx
.
apply
span_EX
in
Hx
.
destruct
Hx
as
[
L
HL
].
apply
(
L
i
).
Defined
.
*
)
End
Dual_basis_Def
.
End
Dual_basis_Def
.
...
@@ -1727,22 +1736,22 @@ Lemma bidual_pt_eval : forall (x : PE_ms), bidual (fun f => val f x).
...
@@ -1727,22 +1736,22 @@ Lemma bidual_pt_eval : forall (x : PE_ms), bidual (fun f => val f x).
Proof
.
intros
x
;
apply
dual_lin_map_equiv
;
easy
.
Qed
.
Proof
.
intros
x
;
apply
dual_lin_map_equiv
;
easy
.
Qed
.
(
*
This
is
the
natural
isomorphism
between
a
subspace
and
its
bidual
.
*
)
(
*
This
is
the
natural
isomorphism
between
a
subspace
and
its
bidual
.
*
)
Definition
bidual_
nat_
isom
(
x
:
PE_ms
)
:
PE
'
'_
ms
:=
Definition
bidual_isom
(
x
:
PE_ms
)
:
PE
'
'_
ms
:=
mk_sub_ms
(
bidual_pt_eval
x
).
mk_sub_ms
(
bidual_pt_eval
x
).
Lemma
bidual_
nat_
isom_correct
:
Lemma
bidual_isom_correct
:
forall
(
x
:
PE_ms
)
(
f
:
PE
'_
ms
),
val
(
bidual_
nat_
isom
x
)
f
=
val
f
x
.
forall
(
x
:
PE_ms
)
(
f
:
PE
'_
ms
),
val
(
bidual_isom
x
)
f
=
val
f
x
.
Proof
.
easy
.
Qed
.
Proof
.
easy
.
Qed
.
Lemma
bidual_
nat_
isom_lin_map
:
is_linear_mapping
bidual_
nat_
isom
.
Lemma
bidual_isom_lin_map
:
is_linear_mapping
bidual_isom
.
Proof
.
Proof
.
Admitted
.
Admitted
.
Lemma
bidual_
nat_
isom_inj
:
injective
bidual_
nat_
isom
.
Lemma
bidual_isom_inj
:
injective
bidual_isom
.
Proof
.
Proof
.
Admitted
.
Admitted
.
Lemma
bidual_
nat_
isom_bij
:
bijective
bidual_
nat_
isom
.
Lemma
bidual_isom_bij
:
bijective
bidual_isom
.
Proof
.
Proof
.
Admitted
.
Admitted
.
...
@@ -1774,7 +1783,7 @@ Let B'' := dual_basis HPE' HB'.
...
@@ -1774,7 +1783,7 @@ Let B'' := dual_basis HPE' HB'.
Let
HB
''
:=
dual_basis_is_basis
HPE
'
HB
'
.
Let
HB
''
:=
dual_basis_is_basis
HPE
'
HB
'
.
Definition
predual_basis
(
i
:
'
I_n
)
:
E
:=
Definition
predual_basis
(
i
:
'
I_n
)
:
E
:=
val
(
f_inv
(
bidual_
nat_
isom_bij
HPE
)
val
(
f_inv
(
bidual_isom_bij
HPE
)
(
mk_sub_ms
(
is_basis_inclF
HB
''
i
))).
(
mk_sub_ms
(
is_basis_inclF
HB
''
i
))).
(
*
(
*
...
...
This diff is collapsed.
Click to expand it.
Preview
0%
Loading
Try again
or
attach a new file
.
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Save comment
Cancel
Please
register
or
sign in
to comment