summaryrefslogtreecommitdiff
path: root/test-suite/success/paralleltac.v
blob: d25fd32a1324a2118797b5e7236cc7c3f863cb6f (plain)
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
Lemma test_nofail_like_all1 :
  True /\ False.
Proof.
split.
all: trivial.
Admitted.

Lemma test_nofail_like_all2 :
  True /\ False.
Proof.
split.
par: trivial.
Admitted.

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 (assert (fib n = S (fib n)) by reflexivity).
(* Tune that depending on your PC *)
Let time := 18.

Axiom P : nat -> Prop.
Axiom P_triv : Type -> forall x, P x.
Ltac solve_P :=
  match goal with |- P (S ?X) =>
    sleep time; exact (P_triv Type _)
  end.

Lemma test_old x : P (S x) /\ P (S x) /\ P (S x) /\ P (S x).
Proof.
repeat split.
idtac "T1: linear".
Time all: solve [solve_P].
Qed.

Lemma test_ok x : P (S x) /\ P (S x) /\ P (S x) /\ P (S x).
Proof.
repeat split.
idtac "T2: parallel".
Time par: solve [solve_P].
Qed.

Lemma test_fail x : P (S x) /\ P x /\ P (S x) /\ P (S x).
Proof.
repeat split.
idtac "T3: linear failure".
Fail Time all: solve solve_P.
all: solve [apply (P_triv Type)].
Qed.

Lemma test_fail2 x : P (S x) /\ P x /\ P (S x) /\ P (S x).
Proof.
repeat split.
idtac "T4: parallel failure".
Fail Time par: solve [solve_P].
all: solve [apply (P_triv Type)].
Qed.