summaryrefslogtreecommitdiff
path: root/test-suite/ssr/havesuff.v
blob: aa1f71879eb703b56e58b4b7d5f146412b4f9d64 (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
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

(* (c) Copyright 2006-2016 Microsoft Corporation and Inria.                  *)

Require Import ssreflect.


Variables P G : Prop.

Lemma test1 : (P -> G) -> P -> G.
Proof.
move=> pg p.
have suff {pg} H : P.
  match goal with |- P -> G => move=> _; exact: pg p | _ => fail end.
match goal with H : P -> G |- G => exact: H p | _ => fail end.
Qed.

Lemma test2 : (P -> G) -> P -> G.
Proof.
move=> pg p.
have suffices {pg} H : P.
  match goal with |- P -> G => move=> _; exact: pg p | _ => fail end.
match goal with H : P -> G |- G => exact: H p | _ => fail end.
Qed.

Lemma test3 : (P -> G) -> P -> G.
Proof.
move=> pg p.
suff have {pg} H : P.
  match goal with H : P |- G => exact: pg H | _ => fail end.
match goal with |- (P -> G) -> G => move=> H; exact: H p | _ => fail end.
Qed.

Lemma test4 : (P -> G) -> P -> G.
Proof.
move=> pg p.
suffices have {pg} H: P.
  match goal with H : P |- G => exact: pg H | _ => fail end.
match goal with |- (P -> G) -> G => move=> H; exact: H p | _ => fail end.
Qed.

(*
Lemma test5 : (P -> G) -> P -> G.
Proof.
move=> pg p.
suff have {pg} H : P := pg H.
match goal with |- (P -> G) -> G => move=> H; exact: H p | _ => fail end.
Qed.
*)

(*
Lemma test6 : (P -> G) -> P -> G.
Proof.
move=> pg p.
suff have {pg} H := pg H.
match goal with |- (P -> G) -> G => move=> H; exact: H p | _ => fail end.
Qed.
*)

Lemma test7 : (P -> G) -> P -> G.
Proof.
move=> pg p.
have suff {pg} H : P := pg.
match goal with H : P -> G |- G => exact: H p | _ => fail end.
Qed.

Lemma test8 : (P -> G) -> P -> G.
Proof.
move=> pg p.
have suff {pg} H := pg.
match goal with H : P -> G |- G => exact: H p | _ => fail end.
Qed.

Goal forall x y : bool, x = y -> x = y.
move=> x y E.
by have {x E} -> : x = y by [].
Qed.