Skip to content

Commit

Permalink
Merge pull request #770 from MetaCoq/fix-float-ops-8.14
Browse files Browse the repository at this point in the history
Fix float ops
  • Loading branch information
mattam82 authored Oct 11, 2022
2 parents 755c691 + 00ca6c9 commit 80aa3a0
Showing 1 changed file with 17 additions and 39 deletions.
56 changes: 17 additions & 39 deletions pcuic/theories/utils/PCUICPrimitive.v
Original file line number Diff line number Diff line change
Expand Up @@ -53,39 +53,8 @@ Proof.
now destruct (uip p q).
Qed.

#[global]
Instance reflect_eq_Z : ReflectEq Z := EqDec_ReflectEq _.

Local Obligation Tactic := idtac.
#[program]
#[global]
Instance reflect_eq_uint63 : ReflectEq uint63_model :=
{ eqb x y := Z.eqb (proj1_sig x) (proj1_sig y) }.
Next Obligation.
cbn -[eqb].
intros x y.
elim: Z.eqb_spec. constructor.
now apply exist_irrel_eq.
intros neq; constructor => H'; apply neq; now subst x.
Qed.

#[global]
Instance reflect_eq_spec_float : ReflectEq SpecFloat.spec_float := EqDec_ReflectEq _.

(* #[program]
#[global]
Instance reflect_eq_float64 : ReflectEq float64_model :=
{ eqb x y := eqb (proj1_sig x) (proj1_sig y) }.
Next Obligation.
cbn -[eqb].
intros x y.
elim: eqb_spec. constructor.
now apply exist_irrel_eq.
intros neq; constructor => H'; apply neq; now subst x.
Qed. *)

(** Propositional UIP is needed below *)
Set Equations With UIP.

Equations eqb_prim_model {term} {t : prim_tag} (x y : prim_model term t) : bool :=
| primIntModel x, primIntModel y := ReflectEq.eqb x y
Expand All @@ -101,20 +70,29 @@ Next Obligation.
Qed.

#[global]
Instance prim_model_eqdec {term} (*e : EqDec term*) : forall p : prim_tag, EqDec (prim_model term p) := _.
Instance prim_model_eqdec {term} : forall p : prim_tag, EqDec (prim_model term p) := _.

Equations eqb_prim_val {term} (x y : prim_val term) : bool :=
| (primInt; i), (primInt; i') := ReflectEq.eqb i i'
| (primFloat; f), (primFloat; f') := ReflectEq.eqb f f'
| x, y := false.

#[global]
Instance prim_tag_model_eqdec term : EqDec (prim_val term).
Proof. eqdec_proof. Defined.
#[global, program]
Instance prim_val_reflect_eq {term} : ReflectEq (prim_val term) :=
{| ReflectEq.eqb := eqb_prim_val |}.
Next Obligation.
intros. funelim (eqb_prim_val x y); simp eqb_prim_val.
case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H. cbn in n. auto.
constructor. intros H; noconf H; auto.
constructor. intros H; noconf H; auto.
case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H; auto.
Qed.

#[global]
Instance prim_val_reflect_eq term : ReflectEq (prim_val term) := EqDec_ReflectEq _.
Instance prim_tag_model_eqdec term : EqDec (prim_val term) := _.

(** Printing *)

Definition string_of_float64_model (f : float64_model) :=
"<>".

Definition string_of_prim {term} (soft : term -> string) (p : prim_val term) : string :=
match p.π2 return string with
| primIntModel f => "(int: " ^ string_of_prim_int f ^ ")"
Expand Down

0 comments on commit 80aa3a0

Please sign in to comment.