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
|
Require Import ZArith.
Require Import Lra.
Require Import Reals.
Goal (1 / (1 - 1) = 0)%R.
Fail lra. (* division by zero *)
Abort.
Goal (0 / (1 - 1) = 0)%R.
lra. (* 0 * x = 0 *)
Qed.
Goal (10 ^ 2 = 100)%R.
lra. (* pow is reified as a constant *)
Qed.
Goal (2 / (1/2) ^ 2 = 8)%R.
lra. (* pow is reified as a constant *)
Qed.
Goal ( IZR (Z.sqrt 4) = 2)%R.
Proof.
Fail lra.
Abort.
Require Import DeclConstant.
#[export] Instance Dsqrt : DeclaredConstant Z.sqrt := {}.
Goal ( IZR (Z.sqrt 4) = 2)%R.
Proof.
lra.
Qed.
Require Import QArith.
Require Import Qreals.
Goal (Q2R (1 # 2) = 1/2)%R.
Proof.
lra.
Qed.
Goal ( 1 ^ (2 + 2) = 1)%R.
Proof.
Fail lra.
Abort.
#[export] Instance Dplus : DeclaredConstant Init.Nat.add := {}.
Goal ( 1 ^ (2 + 2) = 1)%R.
Proof.
lra.
Qed.
Require Import Lia.
Goal ( 1 ^ (2 + 2) = 1)%Z.
Proof.
lia. (* exponent is a constant expr *)
Qed.
Goal (1 / IZR (Z.pow_pos 10 25) = 1 / 10 ^ 25)%R.
Proof.
lra.
Qed.
|