1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79 80 81 82
|
From Coq Require Import Nsatz.
From mathcomp Require Import all_ssreflect ssralg ssrint ssrnum.
From mathcomp Require Import boolp reals constructive_ereal.
(**md**************************************************************************)
(* # 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. *)
(* *)
(* Reference: *)
(* - 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).
Proof. by constructor => [x //|x y //|x y z ->]. Qed.
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)).
Proof.
Defined.
#[global]
Instance Nsatz_realType_Ring : (Ncring.Ring (Ro:=Nsatz_realType_Ring_ops)).
Proof.
constructor => //.
- exact: Nsatz_realType_Setoid_Theory.
- by move=> x y -> x1 y1 ->.
- by move=> x y -> x1 y1 ->.
- by move=> x y -> x1 y1 ->.
- by move=> 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]
Instance Nsatz_realType_Cring: (Cring.Cring (Rr:=Nsatz_realType_Ring)).
Proof. exact: mulrC. Defined.
#[global]
Instance Nsatz_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.
by rewrite mulf_eq0 => /orP[] /eqP->; [left | right].
rewrite -[_ _ Algebra_syntax.zero]/(1 = 0)%R; apply/eqP.
by rewrite (eqr_nat T 1 0).
Defined.
End Nsatz_realType.
Tactic Notation "nsatz" := nsatz_default.
|