FinMatrix.Matrix.ElementType


Require NatExt ZExt QExt QcExt RExt RealFunction Complex.
Require Export Hierarchy.

ElementType

A type with decidable equality and zero element
Module Type ElementType.
  Parameter A : Type.
  Parameter Azero : A.
  Notation "0" := Azero : A_scope.

  Axiom AeqDec : Dec (@eq A).
  #[export] Existing Instance AeqDec.
End ElementType.

Instances

Module ElementTypeNat <: ElementType.
  Export NatExt.
  Definition A : Type := nat.
  Definition Azero : A := 0.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. apply nat_eq_Dec. Defined.
End ElementTypeNat.

Module ElementTypeZ <: ElementType.
  Export ZExt.
  Definition A : Type := Z.
  Definition Azero : A := 0.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. apply Z_eq_Dec. Defined.
End ElementTypeZ.

Module ElementTypeQ <: ElementType.
  Export QExt.
  Definition A : Type := Q.
  Definition Azero : A := 0.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. apply Q_eq_Dec. Defined.
End ElementTypeQ.

Module ElementTypeQc <: ElementType.
  Export QcExt.
  Definition A : Type := Qc.
  Definition Azero : A := 0.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. apply Qc_eq_Dec. Defined.

End ElementTypeQc.

Module ElementTypeR <: ElementType.
  Export RExt.

  Definition A : Type := R.
  Definition Azero : A := 0.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. apply R_eq_Dec. Defined.
End ElementTypeR.

Module ElementTypeC <: ElementType.
  Export Complex.
  Definition A : Type := C.
  Definition Azero : A := 0.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. apply Complex_eq_Dec. Defined.
End ElementTypeC.

Module ElementTypeRFun <: ElementType.
  Export RealFunction.
  Definition A : Type := tpRFun.
  Definition Azero : A := fzero.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. constructor. intros a b. unfold tpRFun in *.
  Admitted.
End ElementTypeRFun.

Module ElementTypeFun (I O : ElementType) <: ElementType.
  Definition A : Type := I.A -> O.A.
  Definition Azero : A := fun _ => O.Azero.
  Hint Unfold A Azero : A.

  Lemma AeqDec : Dec (@eq A).
  Proof. constructor. intros a b. unfold A in *.
  Admitted.
End ElementTypeFun.

Module ElementType_Test.
  Import ElementTypeNat ElementTypeR.
  Module Import ElementTypeFunEx1 :=
    ElementTypeFun ElementTypeNat ElementTypeR.

  Definition f : A := fun i => match i with 0%nat => 1 | 1%nat => 2 | _ => 1 end.
  Definition g : A := fun i => match i with 1%nat => 2 | _ => 1 end.

  Goal f = g.
  Proof. cbv. intros. auto. Qed.
End ElementType_Test.

OrderedElementType

An extended `ElementType` equipped with order relation
Module Type OrderedElementType <: ElementType.
  Include ElementType.

  Parameter Alt Ale : A -> A -> Prop.
  Parameter Altb Aleb : A -> A -> bool.

  Infix "<" := Alt : A_scope.
  Infix "<=" := Ale : A_scope.
  Infix "<?" := Altb : A_scope.
  Infix "<=?" := Aleb : A_scope.

  Axiom Order : Order Alt Ale Altb Aleb.
End OrderedElementType.

Instances

Module OrderedElementTypeNat <: OrderedElementType.
  Include ElementTypeNat.

  Definition Alt := Nat.lt.
  Definition Ale := Nat.le.
  Definition Altb := Nat.ltb.
  Definition Aleb := Nat.leb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply nat_Order. Qed.
End OrderedElementTypeNat.

Module OrderedElementTypeZ <: OrderedElementType.
  Include ElementTypeZ.

  Definition Alt := Z.lt.
  Definition Ale := Z.le.
  Definition Altb := Z.ltb.
  Definition Aleb := Z.leb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply Z_Order. Qed.
End OrderedElementTypeZ.


Module OrderedElementTypeQc <: OrderedElementType.
  Include ElementTypeQc.

  Definition Alt := Qclt.
  Definition Ale := Qcle.
  Definition Altb := Qcltb.
  Definition Aleb := Qcleb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply Qc_Order. Qed.
End OrderedElementTypeQc.

Module OrderedElementTypeR <: OrderedElementType.
  Include ElementTypeR.

  Definition Alt := Rlt.
  Definition Ale := Rle.
  Definition Altb := Rltb.
  Definition Aleb := Rleb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply R_Order. Qed.
End OrderedElementTypeR.

Element with abelian-monoid structure

Type of Element with abelian-monoid structure

Module Type MonoidElementType <: ElementType.
  Include ElementType.
  Open Scope A_scope.

  Parameter Aadd : A -> A -> A.
  Infix "+" := Aadd : A_scope.

  Axiom Aadd_AMonoid : AMonoid Aadd Azero.
  #[export] Existing Instance Aadd_AMonoid.
End MonoidElementType.

Instances


Module MonoidElementTypeNat <: MonoidElementType.
  Include ElementTypeNat.

  Definition Aadd := Nat.add.
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof. intros. repeat constructor; intros; autounfold with A; ring. Qed.
End MonoidElementTypeNat.

Module MonoidElementTypeZ <: MonoidElementType.
  Include ElementTypeZ.

  Definition Aadd := Zplus.
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof. intros. repeat constructor; intros; autounfold with A; ring. Qed.
End MonoidElementTypeZ.

Module MonoidElementTypeQc <: MonoidElementType.
  Include ElementTypeQc.

  Definition Aadd := Qcplus.
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof. intros. repeat constructor; intros; autounfold with A; ring. Qed.
End MonoidElementTypeQc.

Module MonoidElementTypeR <: MonoidElementType.
  Include ElementTypeR.

  Definition Aadd := Rplus.
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof. intros. repeat constructor; intros; autounfold with A; ring. Qed.
End MonoidElementTypeR.

Module MonoidElementTypeC <: MonoidElementType.
  Include ElementTypeC.

  Definition Aadd := Cadd.

Note that, this explicit annotation is must, otherwise, the ring has no effect. (because C and T are different)
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof. intros. repeat constructor; intros; autounfold with A; ring. Qed.
End MonoidElementTypeC.

Module MonoidElementTypeRFun <: MonoidElementType.
  Include ElementTypeRFun.

  Definition Aadd := fadd.
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof. intros. repeat constructor; intros; autounfold with A; ring. Qed.
End MonoidElementTypeRFun.

Module MonoidElementTypeFun (I O : MonoidElementType) <: MonoidElementType.

  Include (ElementTypeFun I O).
  Open Scope A_scope.

  Definition Aadd (f g : A) : A := fun x : I.A => O.Aadd (f x) (g x).
  Hint Unfold Aadd : A.

  Infix "+" := Aadd : A_scope.

  Lemma Aadd_Associative : Associative Aadd.
  Proof.
    intros. constructor; intros; autounfold with A.
    extensionality x. apply O.Aadd_AMonoid.
  Qed.

  Lemma Aadd_Commutative : Commutative Aadd.
  Proof.
    intros. constructor; intros; autounfold with A.
    extensionality x. apply O.Aadd_AMonoid.
  Qed.

  Lemma Aadd_IdentityLeft : IdentityLeft Aadd Azero.
  Proof.
    intros. constructor; intros; autounfold with A.
    extensionality x. apply O.Aadd_AMonoid.
  Qed.

  Lemma Aadd_IdentityRight : IdentityRight Aadd Azero.
  Proof.
    intros. constructor; intros; autounfold with A.
    extensionality x. apply O.Aadd_AMonoid.
  Qed.

  #[export] Instance Aadd_AMonoid : AMonoid Aadd Azero.
  Proof.
    intros. constructor; intros; autounfold with A.
    intros. constructor; intros; autounfold with A.
    constructor. apply Aadd_Associative.
    apply Aadd_IdentityLeft. apply Aadd_IdentityRight.
    constructor. apply Aadd_Associative.
    apply Aadd_Commutative.
    constructor. constructor. apply Aadd_Associative.
    apply Aadd_Commutative.
  Qed.
End MonoidElementTypeFun.

Module MonoidElementTypeTest.
  Import MonoidElementTypeQc.
  Import MonoidElementTypeR.

  Module Import MonoidElementTypeFunEx1 :=
    MonoidElementTypeFun MonoidElementTypeQc MonoidElementTypeR.

  Definition f : A := fun i => 1.
  Definition g : A := fun i => 2.
  Definition h : A := fun i => 3.

  Goal f + g + h = f + (g + h).
  Proof. rewrite associative. auto. Qed.
End MonoidElementTypeTest.

Element with abelian-ring structure

Type of Element with abelian-ring structure

Module Type RingElementType <: MonoidElementType.
  Include MonoidElementType.
  Open Scope A_scope.

  Parameter Aone : A.
  Parameter Amul : A -> A -> A.
  Parameter Aopp : A -> A.

  Notation Asub := (fun x y => Aadd x (Aopp y)).
  Infix "*" := Amul : A_scope.
  Notation "1" := Aone : A_scope.
  Notation "a ²" := (a * a) : A_scope.
  Notation "- a" := (Aopp a) : A_scope.
  Infix "-" := Asub : A_scope.

A Ring structure can be derived from the context
  Axiom ARing : ARing Aadd 0 Aopp Amul 1.

  #[export] Existing Instance ARing.

End RingElementType.

Instances


Module RingElementTypeZ <: RingElementType.
  Include MonoidElementTypeZ.

  Definition Aone : A := 1.
  Definition Aopp := Z.opp.
  Definition Amul := Zmult.
  Hint Unfold Aone Aopp Amul : A.

  Notation Asub := (fun x y => Aadd x (Aopp y)).
  Infix "*" := Amul : A_scope.
  Notation "- a" := (Aopp a) : A_scope.
  Infix "-" := Asub : A_scope.

  #[export] Instance ARing : ARing Aadd Azero Aopp Amul Aone.
  Proof. repeat constructor; autounfold with A; intros; ring. Qed.

End RingElementTypeZ.

Module RingElementTypeQc <: RingElementType.
  Include MonoidElementTypeQc.

  Definition Aone : A := 1.
  Definition Aopp := Qcopp.
  Definition Amul := Qcmult.
  Hint Unfold Aone Aadd Aopp Amul : A.

  Notation Asub := (fun x y => Aadd x (Aopp y)).
  Infix "*" := Amul : A_scope.
  Notation "- a" := (Aopp a) : A_scope.
  Infix "-" := Asub : A_scope.

  #[export] Instance ARing : ARing Aadd Azero Aopp Amul Aone.
  Proof. repeat constructor; autounfold with A; intros; ring. Qed.

  Add Ring Ring_inst : (make_ring_theory ARing).
End RingElementTypeQc.

Module RingElementTypeR <: RingElementType.
  Include MonoidElementTypeR.

  Definition Aone : A := 1.
  Definition Aopp := Ropp.
  Definition Amul := Rmult.
  Hint Unfold Aone Aadd Aopp Amul : A.

  Notation Asub := (fun x y => Aadd x (Aopp y)).
  Infix "*" := Amul : A_scope.
  Notation "- a" := (Aopp a) : A_scope.
  Infix "-" := Asub : A_scope.

  #[export] Instance ARing : ARing Aadd Azero Aopp Amul Aone.
  Proof. repeat constructor; autounfold with A; intros; ring. Qed.

End RingElementTypeR.

Module RingElementTypeC <: RingElementType.
  Include MonoidElementTypeC.

  Definition Aone : A := 1.
  Definition Aopp := Copp.
  Definition Amul := Cmul.
  Hint Unfold Aone Aadd Aopp Amul : A.

  Notation Asub := (fun x y => Aadd x (Aopp y)).
  Infix "*" := Amul : A_scope.
  Notation "- a" := (Aopp a) : A_scope.
  Infix "-" := Asub : A_scope.

  #[export] Instance ARing : ARing Aadd Azero Aopp Amul Aone.
  Proof. repeat constructor; autounfold with A; intros; ring. Qed.

  Add Ring Ring_inst : (make_ring_theory ARing).
End RingElementTypeC.

Module RingElementTypeRFun <: RingElementType.
  Include MonoidElementTypeRFun.

  Definition Aone : A := fone.
  Definition Aopp := fopp.
  Definition Amul := fmul.
  Hint Unfold Aone Aadd Aopp Amul : A.

  Notation Asub := (fun x y => Aadd x (Aopp y)).
  Infix "*" := Amul : A_scope.
  Notation "- a" := (Aopp a) : A_scope.
  Infix "-" := Asub : A_scope.

  #[export] Instance ARing : ARing Aadd Azero Aopp Amul Aone.
  Proof.
    repeat constructor; intros; autounfold with A;
      apply functional_extensionality; intros; cbv; ring.
  Qed.

  Add Ring Ring_inst : (make_ring_theory ARing).
End RingElementTypeRFun.

Module RingElementTypeFun (I O : RingElementType) <: RingElementType.
  Add Ring Ring_thy_inst_o : (make_ring_theory O.ARing).

  Include (MonoidElementTypeFun I O).

  Definition Aone : A := fun _ => O.Aone.
  Definition Aopp (f : A) : A := fun x : I.A => O.Aopp (f x).
  Definition Amul (f g : A) : A := fun x : I.A => O.Amul (f x) (g x).
  Notation Asub := (fun x y => Aadd x (Aopp y)).

  #[export] Instance ARing : ARing Aadd Azero Aopp Amul Aone.
  Proof.
    repeat constructor; autounfold with A; intros;
      apply functional_extensionality; intros; cbv; ring.
  Qed.

  Add Ring Ring_inst : (make_ring_theory ARing).
End RingElementTypeFun.

Module RingElementTypeTest.
  Import RingElementTypeQc.
  Import RingElementTypeR.

  Module Import RingElementTypeFunEx1 :=
    RingElementTypeFun RingElementTypeQc RingElementTypeR.

  Definition f : A := fun i:Qc => (Qc2R i + R1)%R.
  Definition g : A := fun i:Qc => Qc2R (i+1).

  Goal f = g.
  Proof. Abort.
End RingElementTypeTest.

Element with abelian-ring structure and ordered relation

Type of Element with ordered abelian-ring structure

Module Type OrderedRingElementType <: RingElementType <: OrderedElementType.
  Include RingElementType.

  Parameter Alt Ale : A -> A -> Prop.
  Parameter Altb Aleb : A -> A -> bool.

  Infix "<" := Alt : A_scope.
  Infix "<=" := Ale : A_scope.
  Infix "<?" := Altb : A_scope.
  Infix "<=?" := Aleb : A_scope.

  Axiom Order : Order Alt Ale Altb Aleb.
  Axiom OrderedARing : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  #[export] Existing Instance OrderedARing.

  Notation "| a |" := (@Aabs _ 0 Aopp Aleb a) : A_scope.

End OrderedRingElementType.

Instances


Module OrderedRingElementTypeZ <: OrderedRingElementType.
  Include RingElementTypeZ.

  Definition Ale := Z.le.
  Definition Alt := Z.lt.
  Definition Altb := Z.ltb.
  Definition Aleb := Z.leb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply OrderedElementTypeZ.Order. Qed.

  #[export] Instance OrderedARing
    : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  Proof.
    constructor. apply ARing. apply Order.
    - intros; autounfold with A in *. lia.
    - intros; autounfold with A in *. apply Zmult_lt_compat_r; auto.
  Qed.
End OrderedRingElementTypeZ.

Module OrderedRingElementTypeQc <: OrderedRingElementType.
  Include RingElementTypeQc.

  Definition Ale := Qcle.
  Definition Alt := Qclt.
  Definition Altb := Qcltb.
  Definition Aleb := Qcleb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply OrderedElementTypeQc.Order. Qed.

  #[export] Instance OrderedARing
    : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  Proof.
    constructor. apply ARing. apply Order.
    - intros; autounfold with A in *. apply Qc_lt_add_compat_r; auto.
    - intros; autounfold with A in *. apply Qc_lt_mul_compat_r; auto.
  Qed.
End OrderedRingElementTypeQc.

Module OrderedRingElementTypeR <: OrderedRingElementType.
  Include RingElementTypeR.

  Definition Ale := Rle.
  Definition Alt := Rlt.
  Definition Altb := Rltb.
  Definition Aleb := Rleb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply OrderedElementTypeR.Order. Qed.

  #[export] Instance OrderedARing
    : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  Proof.
    constructor. apply ARing. apply Order.
    - intros; autounfold with A in *. lra.
    - intros; autounfold with A in *. apply Rmult_lt_compat_r; auto.
  Qed.

  Notation "| a |" := (@Aabs _ 0 Aopp Aleb a) : A_scope.

End OrderedRingElementTypeR.

Element with field structure

Type of Element with field structure


Module Type FieldElementType <: RingElementType.
  Include RingElementType.
  Open Scope A_scope.

  Parameter Ainv : A -> A.

  Notation Adiv := (fun x y => Amul x (Ainv y)).
  Notation "/ a" := (Ainv a) : A_scope.
  Infix "/" := Adiv : A_scope.

1 <> 0.
  Axiom Aone_neq_Azero : Aone <> Azero.

  Axiom Field : Field Aadd Azero Aopp Amul Aone Ainv.

  #[export] Existing Instance Field.
End FieldElementType.

Instances


Module FieldElementTypeQc <: FieldElementType.
  Include RingElementTypeQc.

  Definition Ainv := Qcinv.
  Hint Unfold Ainv : A.

  Notation Adiv := (fun x y => Amul x (Ainv y)).

  Lemma Aone_neq_Azero : Aone <> Azero.
  Proof. intro. cbv in *. inv H. Qed.

  #[export] Instance Field : Field Aadd Azero Aopp Amul Aone Ainv.
  Proof.
    constructor. apply ARing.
    intros. autounfold with A. field. auto.
    apply Aone_neq_Azero.
  Qed.

  Add Field Field_inst : (make_field_theory Field).
End FieldElementTypeQc.

Module FieldElementTypeR <: FieldElementType.
  Include RingElementTypeR.

  Definition Ainv := Rinv.
  Hint Unfold Ainv : A.

  Notation Adiv := (fun x y => Amul x (Ainv y)).

  Lemma Aone_neq_Azero : Aone <> Azero.
  Proof. cbv in *. auto with real. Qed.

  #[export] Instance Field : Field Aadd Azero Aopp Amul Aone Ainv.
  Proof.
    constructor. apply ARing. intros.
    autounfold with A. field. auto.
    apply Aone_neq_Azero.
  Qed.

End FieldElementTypeR.

Module FieldElementTypeC <: FieldElementType.
  Include RingElementTypeC.

  Definition Ainv := Cinv.
  Hint Unfold Ainv : A.

  Notation Adiv := (fun x y => Amul x (Ainv y)).

  Lemma Aone_neq_Azero : Aone <> Azero.
  Proof. cbv in *. auto with complex. Qed.

  #[export] Instance Field : Field Aadd Azero Aopp Amul Aone Ainv.
  Proof.
    constructor. apply ARing. intros.
    autounfold with A. field. auto.
    apply Aone_neq_Azero.
  Qed.

  Add Field Field_inst : (make_field_theory Field).
End FieldElementTypeC.







Module FieldElementTypeTest.
  Import FunctionalExtensionality.
  Import FieldElementTypeQc.
  Import FieldElementTypeR.

End FieldElementTypeTest.

Element with field structure and ordered relation

Type of Element with ordered field structure

Module Type OrderedFieldElementType <: FieldElementType <: OrderedRingElementType.
  Include FieldElementType.

  Parameter Alt Ale : A -> A -> Prop.
  Parameter Altb Aleb : A -> A -> bool.

  Infix "<" := Alt : A_scope.
  Infix "<=" := Ale : A_scope.
  Infix "<?" := Altb : A_scope.
  Infix "<=?" := Aleb : A_scope.

  Axiom Order : Order Alt Ale Altb Aleb.

  Axiom OrderedARing : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  #[export] Existing Instance OrderedARing.

  Axiom OrderedAField : OrderedField Aadd Azero Aopp Amul Aone Ainv Alt Ale Altb Aleb.
  #[export] Existing Instance OrderedAField.

  Notation "| a |" := (@Aabs _ 0 Aopp Aleb a) : A_scope.
End OrderedFieldElementType.

Instances


Module OrderedFieldElementTypeQc <: OrderedFieldElementType.
  Include FieldElementTypeQc.

  Definition Ale := Qcle.
  Definition Alt := Qclt.
  Definition Altb := Qcltb.
  Definition Aleb := Qcleb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply OrderedElementTypeQc.Order. Qed.

  #[export] Instance OrderedARing
    : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  Proof. apply OrderedRingElementTypeQc.OrderedARing. Qed.

  #[export] Instance OrderedAField
    : OrderedField Aadd Azero Aopp Amul Aone Ainv Alt Ale Altb Aleb.
  Proof. constructor. apply Field. apply OrderedRingElementTypeQc.OrderedARing. Qed.

End OrderedFieldElementTypeQc.

Module OrderedFieldElementTypeR <: OrderedFieldElementType.
  Include FieldElementTypeR.

  Definition Ale := Rle.
  Definition Alt := Rlt.
  Definition Altb := Rltb.
  Definition Aleb := Rleb.
  Hint Unfold Ale Alt Altb Aleb : A.

  #[export] Instance Order : Order Alt Ale Altb Aleb.
  Proof. apply OrderedElementTypeR.Order. Qed.

  #[export] Instance OrderedARing
    : OrderedARing Aadd Azero Aopp Amul Aone Alt Ale Altb Aleb.
  Proof. apply OrderedRingElementTypeR.OrderedARing. Qed.

  #[export] Instance OrderedAField
    : OrderedField Aadd Azero Aopp Amul Aone Ainv Alt Ale Altb Aleb.
  Proof. constructor. apply Field. apply OrderedRingElementTypeR.OrderedARing. Qed.

  Notation "| a |" := (@Aabs _ 0 Aopp Aleb a) : A_scope.

End OrderedFieldElementTypeR.

Element with field structure and ordered relation and normed

Type of Element with normed ordered field structure

Module Type NormedOrderedFieldElementType <: OrderedFieldElementType.
  Include OrderedFieldElementType.
  Export RExt RealFunction.

  Parameter a2r : A -> R.
  Axiom ConvertToR : ConvertToR Aadd Azero Aopp Amul Aone Ainv Alt Ale Altb Aleb a2r.
  #[export] Existing Instance ConvertToR.
End NormedOrderedFieldElementType.

Instances