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
|
(* Some boilerplate *)
Fixpoint fib n := match n with
| O => 1
| S m => match m with
| O => 1
| S o => fib o + fib m end end.
Ltac sleep n :=
try (cut (fib n = S (fib n)); reflexivity).
(* Tune that depending on your PC *)
Let time := 10.
(* Beginning of demo *)
Section Demo.
Variable i : True.
Lemma a (n : nat) : nat.
Proof using.
revert n.
fix f 1.
apply f.
Qed.
Lemma b : True.
Proof using i.
sleep time.
idtac.
sleep time.
(* Here we use "a" *)
exact I.
Qed.
Lemma work_here : True /\ True.
Proof using i.
(* Jump directly here, Coq reacts immediately *)
split.
exact b. (* We can use the lemmas above *)
exact I.
Qed.
End Demo.
From Coq Require Import Program.Tactics.
Obligation Tactic := idtac.
Program Definition foo : nat -> nat * nat :=
fix f (n : nat) := (0,_).
Next Obligation.
intros f n; apply (f n).
Qed.
|