From 7cfc4e5146be5666419451bdd516f1f3f264d24a Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Sun, 25 Jan 2015 14:42:51 +0100 Subject: Imported Upstream version 8.5~beta1+dfsg --- test-suite/bugs/closed/3652.v | 101 ++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 101 insertions(+) create mode 100644 test-suite/bugs/closed/3652.v (limited to 'test-suite/bugs/closed/3652.v') diff --git a/test-suite/bugs/closed/3652.v b/test-suite/bugs/closed/3652.v new file mode 100644 index 00000000..86e06137 --- /dev/null +++ b/test-suite/bugs/closed/3652.v @@ -0,0 +1,101 @@ +Require Setoid. +Require ZArith. +Import ZArith. + +Inductive Erasable(A : Set) : Prop := + erasable: A -> Erasable A. + +Arguments erasable [A] _. + +Hint Constructors Erasable. + +Scheme Erasable_elim := Induction for Erasable Sort Prop. + +Notation "## T" := (Erasable T) (at level 1, format "## T") : Erasable_scope. +Notation "# x" := (erasable x) (at level 1, format "# x") : Erasable_scope. +Open Scope Erasable_scope. + +Axiom Erasable_inj : forall {A : Set}{a b : A}, #a=#b -> a=b. + +Lemma Erasable_rw : forall (A: Set)(a b : A), (#a=#b) <-> (a=b). +Proof. + intros A a b. + split. + - apply Erasable_inj. + - congruence. +Qed. + +Open Scope Z_scope. +Opaque Z.mul. + +Infix "^" := Zpower_nat : Z_scope. + +Notation "f ; v <- x" := (let (v) := x in f) + (at level 199, left associativity) : Erasable_scope. +Notation "f ; < v" := (f ; v <- v) + (at level 199, left associativity) : Erasable_scope. +Notation "f |# v <- x" := (#f ; v <- x) + (at level 199, left associativity) : Erasable_scope. +Notation "f |# < v" := (#f ; < v) + (at level 199, left associativity) : Erasable_scope. + +Ltac name_evars id := + repeat match goal with |- context[?V] => + is_evar V; let H := fresh id in set (H:=V) in * end. + +Lemma Twoto0 : 2^0 = 1. +Proof. compute. reflexivity. Qed. + +Ltac ring_simplify' := rewrite ?Twoto0; ring_simplify. + +Definition mp2a1s(x : Z)(n : nat) := x * 2^n + (2^n-1). + +Hint Unfold mp2a1s. + +Definition zotval(n1s : nat)(is2 : bool)(next_value : Z) : Z := + 2 * mp2a1s next_value n1s + if is2 then 2 else 0. + +Inductive zot'(eis2 : ##bool)(value : ##Z) : Set := +| Zot'(is2 : bool) + (iseq : eis2=#is2) + {next_is2 : ##bool} + (ok : is2=true -> next_is2=#false) + {next_value : ##Z} + (n1s : nat) + (veq : value = (zotval n1s is2 next_value |#