Library mathcomp.analysis.nsatz_realtype
Require Import Nsatz.
From mathcomp Require Import all_ssreflect ssralg ssrint ssrnum.
From mathcomp.classical Require Import boolp.
Require Import reals ereal.
From mathcomp Require Import all_ssreflect ssralg ssrint ssrnum.
From mathcomp.classical Require Import boolp.
Require Import reals ereal.
nsatz for realType
This file registers the ring corresponding to the MathComp-Analysis type
realType to the tactic nsatz of Coq. This enables some automation used for
example in the file trigo.v.
ref: https://coq.inria.fr/refman/addendum/nsatz.html
Import GRing.Theory Num.Theory.
Local Open Scope ring_scope.
Section Nsatz_realType.
Variable T : realType.
Lemma Nsatz_realType_Setoid_Theory : Setoid.Setoid_Theory T (@eq T).
Definition Nsatz_realType0 := (0%:R : T).
Definition Nsatz_realType1 := (1%:R : T).
Definition Nsatz_realType_add (x y : T) := (x + y)%R.
Definition Nsatz_realType_mul (x y : T) := (x × y)%R.
Definition Nsatz_realType_sub (x y : T) := (x - y)%R.
Definition Nsatz_realType_opp (x : T) := (- x)%R.
#[global]
Instance Nsatz_realType_Ring_ops:
(@Ncring.Ring_ops T Nsatz_realType0 Nsatz_realType1
Nsatz_realType_add
Nsatz_realType_mul
Nsatz_realType_sub
Nsatz_realType_opp (@eq T)).
Defined.
#[global]
Instance Nsatz_realType_Ring : (Ncring.Ring (Ro:=Nsatz_realType_Ring_ops)).
#[global]
Instance Nsatz_realType_Cring: (Cring.Cring (Rr:=Nsatz_realType_Ring)).
#[global]
Instance Nsatz_realType_Integral_domain :
(Integral_domain.Integral_domain (Rcr:=Nsatz_realType_Cring)).
End Nsatz_realType.
Tactic Notation "nsatz" := nsatz_default.