Built with Alectryon, running Coq+SerAPI v8.15.0+0.15.0. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
Require Import Nsatz.
Notation"[ predI _ & _ ]" was already used in scope
fun_scope. [notation-overridden,parsing]
Notation"[ predU _ & _ ]" was already used in scope
fun_scope. [notation-overridden,parsing]
Notation"[ predD _ & _ ]" was already used in scope
fun_scope. [notation-overridden,parsing]
Notation"[ predC _ ]" was already used in scope
fun_scope. [notation-overridden,parsing]
Notation"[ preim _ of _ ]" was already used in scope
fun_scope. [notation-overridden,parsing]
Notation"_ + _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ - _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ <= _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ < _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ >= _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ > _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ <= _ <= _" was already used in scope
nat_scope. [notation-overridden,parsing]
Notation"_ < _ <= _" was already used in scope
nat_scope. [notation-overridden,parsing]
Notation"_ <= _ < _" was already used in scope
nat_scope. [notation-overridden,parsing]
Notation"_ < _ < _" was already used in scope
nat_scope. [notation-overridden,parsing]
Notation"_ * _" was already used in scope nat_scope.
[notation-overridden,parsing]
Notation"_ ^ _" was already used in scope nat_scope.
[notation-overridden,parsing]
New coercion path [GRing.subring_closedM;
GRing.smulr_closedN] : GRing.subring_closed >-> GRing.oppr_closed is ambiguous with existing
[GRing.subring_closedB; GRing.zmod_closedN] : GRing.subring_closed >-> GRing.oppr_closed.
[ambiguous-paths,typechecker]
New coercion path [GRing.subring_closed_semi;
GRing.semiring_closedM] : GRing.subring_closed >-> GRing.mulr_closed is ambiguous with existing
[GRing.subring_closedM; GRing.smulr_closedM] : GRing.subring_closed >-> GRing.mulr_closed.
New coercion path [GRing.subring_closed_semi;
GRing.semiring_closedD] : GRing.subring_closed >-> GRing.addr_closed is ambiguous with existing
[GRing.subring_closedB; GRing.zmod_closedD] : GRing.subring_closed >-> GRing.addr_closed.
[ambiguous-paths,typechecker]
New coercion path [GRing.sdivr_closed_div;
GRing.divr_closedM] : GRing.sdivr_closed >-> GRing.mulr_closed is ambiguous with existing
[GRing.sdivr_closedM; GRing.smulr_closedM] : GRing.sdivr_closed >-> GRing.mulr_closed.
[ambiguous-paths,typechecker]
New coercion path [GRing.subalg_closedBM;
GRing.subring_closedB] : GRing.subalg_closed >-> GRing.zmod_closed is ambiguous with existing
[GRing.subalg_closedZ; GRing.submod_closedB] : GRing.subalg_closed >-> GRing.zmod_closed.
[ambiguous-paths,typechecker]
New coercion path [GRing.divring_closed_div;
GRing.sdivr_closedM] : GRing.divring_closed >-> GRing.smulr_closed is ambiguous with existing
[GRing.divring_closedBM; GRing.subring_closedM] : GRing.divring_closed >-> GRing.smulr_closed.
[ambiguous-paths,typechecker]
New coercion path [GRing.divalg_closedBdiv;
GRing.divring_closedBM] : GRing.divalg_closed >-> GRing.subring_closed is ambiguous with existing
[GRing.divalg_closedZ; GRing.subalg_closedBM] : GRing.divalg_closed >-> GRing.subring_closed.
[ambiguous-paths,typechecker]
New coercion path [GRing.Pred.subring_smul;
GRing.Pred.smul_mul] : GRing.Pred.subring >-> GRing.Pred.mul is ambiguous with existing
[GRing.Pred.subring_semi; GRing.Pred.semiring_mul] : GRing.Pred.subring >-> GRing.Pred.mul.
[ambiguous-paths,typechecker]
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.
LocalOpen Scope ring_scope.
SectionNsatz_realType.
VariableT : realType.
LemmaNsatz_realType_Setoid_Theory : Setoid.Setoid_Theory T (@eq T).
Proof. byconstructor => [x //|x y //|x y z ->]. Qed.
DefinitionNsatz_realType0 := (0%:R : T).
DefinitionNsatz_realType1 := (1%:R : T).
DefinitionNsatz_realType_add (xy : T) := (x + y)%R.
DefinitionNsatz_realType_mul (xy : T) := (x * y)%R.
DefinitionNsatz_realType_sub (xy : T) := (x - y)%R.
DefinitionNsatz_realType_opp (x : T) := (- x)%R.
#[global]
InstanceNsatz_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]
InstanceNsatz_realType_Ring : (Ncring.Ring (Ro:=Nsatz_realType_Ring_ops)).
Proof.
constructor => //.
- exact: Nsatz_realType_Setoid_Theory.
- bymove=> x y -> x1 y1 ->.
- bymove=> x y -> x1 y1 ->.
- bymove=> x y -> x1 y1 ->.
- bymove=> x y ->.
- exact: add0r.
- exact: addrC.
- exact: addrA.
- exact: mul1r.
- exact: mulr1.
- exact: mulrA.
- exact: mulrDl.
- move=> x y z; exact: mulrDr.
- exact: subrr.
Defined.
#[global]
InstanceNsatz_realType_Cring: (Cring.Cring (Rr:=Nsatz_realType_Ring)).
Proof. exact: mulrC. Defined.
#[global]
InstanceNsatz_realType_Integral_domain :
(Integral_domain.Integral_domain (Rcr:=Nsatz_realType_Cring)).
Proof.
constructor.
move=> x y.
rewrite -[_ _ Algebra_syntax.zero]/(x * y = 0)%R => /eqP.
byrewrite mulf_eq0 => /orP[] /eqP->; [left | right].
rewrite -[_ _ Algebra_syntax.zero]/(1 = 0)%R; apply/eqP.
byrewrite (eqr_nat T 10).
Defined.
EndNsatz_realType.
Tactic Notation"nsatz" := nsatz_default.