Skip to content

Commit

Permalink
Add instances & lemmas for regular algebras
Browse files Browse the repository at this point in the history
  • Loading branch information
Georges Gonthier committed Dec 4, 2015
1 parent bd8cf1b commit efc830f
Show file tree
Hide file tree
Showing 3 changed files with 16 additions and 0 deletions.
5 changes: 5 additions & 0 deletions mathcomp/field/falgebra.v
Original file line number Diff line number Diff line change
Expand Up @@ -214,6 +214,11 @@ Notation "1" := (vline 1) : vspace_scope.

Canonical matrix_FalgType (K : fieldType) n := [FalgType K of 'M[K]_n.+1].

Canonical regular_FalgType (R : comUnitRingType) := [FalgType R of R^o].

Lemma regular_fullv (K : fieldType) : (fullv = 1 :> {vspace K^o})%VS.
Proof. by apply/esym/eqP; rewrite eqEdim subvf dim_vline oner_eq0 dimvf. Qed.

Section Proper.

Variables (R : ringType) (aT : FalgType R).
Expand Down
2 changes: 2 additions & 0 deletions mathcomp/field/fieldext.v
Original file line number Diff line number Diff line change
Expand Up @@ -265,6 +265,8 @@ End Exports.
End FieldExt.
Export FieldExt.Exports.

Canonical regular_fieldExtType (F : fieldType) := [fieldExtType F of F^o for F].

Section FieldExtTheory.

Variables (F0 : fieldType) (L : fieldExtType F0).
Expand Down
9 changes: 9 additions & 0 deletions mathcomp/field/galois.v
Original file line number Diff line number Diff line change
Expand Up @@ -456,6 +456,15 @@ apply/seqv_sub_adjoin/imageP; rewrite (tnth_nth 0) /in_mem/=.
by exists (i, widen_ord (sz_r i) j) => /=.
Qed.

Fact regular_splittingAxiom F : SplittingField.axiom (regular_fieldExtType F).
Proof.
exists 1; first exact: rpred1.
by exists [::]; [rewrite big_nil eqpxx | rewrite Fadjoin_nil regular_fullv].
Qed.

Canonical regular_splittingFieldType (F : fieldType) :=
SplittingFieldType F F^o (regular_splittingAxiom F).

Section SplittingFieldTheory.

Variables (F : fieldType) (L : splittingFieldType F).
Expand Down

0 comments on commit efc830f

Please sign in to comment.