From 22dfbff296cf03b6fab2bcec4eb5f9cf6ee8368c Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 4 Nov 2016 15:55:52 +0100 Subject: Fix #3441 Use pf_get_type_of to avoid blowup ... in pose proof of large proof terms --- test-suite/bugs/closed/3441.v | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) create mode 100644 test-suite/bugs/closed/3441.v (limited to 'test-suite/bugs/closed/3441.v') diff --git a/test-suite/bugs/closed/3441.v b/test-suite/bugs/closed/3441.v new file mode 100644 index 000000000..50d297807 --- /dev/null +++ b/test-suite/bugs/closed/3441.v @@ -0,0 +1,23 @@ +Axiom f : nat -> nat -> nat. +Fixpoint do_n (n : nat) (k : nat) := + match n with + | 0 => k + | S n' => do_n n' (f k k) + end. + +Notation big := (_ = _). +Axiom k : nat. +Goal True. +Timeout 1 let H := fresh "H" in + let x := constr:(let n := 17 in do_n n = do_n n) in + let y := (eval lazy in x) in + pose proof y as H. (* Finished transaction in 1.102 secs (1.084u,0.016s) (successful) *) +Timeout 1 let H := fresh "H" in + let x := constr:(let n := 17 in do_n n = do_n n) in + let y := (eval lazy in x) in + pose y as H; clearbody H. (* Finished transaction in 0.412 secs (0.412u,0.s) (successful) *) + +Timeout 1 Time let H := fresh "H" in + let x := constr:(let n := 17 in do_n n = do_n n) in + let y := (eval lazy in x) in + assert (H := y). (* Finished transaction in 1.19 secs (1.164u,0.024s) (successful) *) \ No newline at end of file -- cgit v1.2.3