File: std_tactics.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (142 lines) | stat: -rw-r--r-- 3,117 bytes parent folder | download | duplicates (2)
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
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
Require Import Coq.Setoids.Setoid.
Require Import Ltac2.Ltac2.

Axiom f: nat -> nat.
Definition g := f.

Axiom Foo1: nat -> Prop.
Axiom Foo2: nat -> Prop.
Axiom Impl: forall n: nat, Foo1 (f n) -> Foo2 (f n).

Create HintDb foo discriminated.
#[export] Hint Constants Opaque : foo.
#[export] Hint Resolve Impl : foo.

Goal forall x, Foo1 (f x) -> Foo2 (g x).
Proof.
  auto with foo.
  #[export] Hint Transparent g : foo.
  auto with foo.
Qed.

Goal forall (x: nat), exists y, f x = g y.
Proof.
  intros.
  eexists.
  unify f g.
  lazy_match! goal with
  | [ |- ?a ?b = ?rhs ] => unify ($a $b) $rhs
  end.
Abort.

Goal forall (x: nat), exists y, f x = g y.
Proof.
  intros.
  eexists.
  Unification.unify TransparentState.full 'f 'g.
  lazy_match! goal with
  | [ |- ?a ?b = ?rhs ] => Unification.unify_with_full_ts '($a $b) rhs
  end.
Abort.

Goal True.
Proof.
  Fail Unification.unify TransparentState.empty '(1 + 1) '2.
  Unification.unify TransparentState.full '(1 + 1) '2.
  Unification.unify (TransparentState.current ()) '(1 + 1) '2.
  Opaque Nat.add.
  Fail Unification.unify (TransparentState.current ()) '(1 + 1) '2.
  Succeed Unification.unify TransparentState.full '(1 + 1) '2.
  exact I.
Qed.


(* Test that by clause of assert doesn't eat all semicolons:
   https://github.com/coq/coq/issues/17491 *)
Goal forall (a: nat), a = a.
Proof.
  intros.
  assert (a = a) by Std.reflexivity ();
  assumption.
Qed.

(* Test that notations in by clause still work: *)
Goal forall (a: nat), a = a.
Proof.
  intros.
  assert (a = a) by exact eq_refl;
  assumption.
Qed.

Goal forall x, (forall (y : unit), y = x) -> forall (x: unit), x = x.
Proof.
  intros x H y.
  rewrite -> H at 1 2.
  reflexivity.
Qed.

Goal forall x, (forall (y : unit), x = y) -> forall (x: unit), x = x.
Proof.
  intros x H y.
  rewrite <- H at 1 2.
  reflexivity.
Qed.

Goal forall x, (forall y : unit, y = x) -> forall y : unit, y = y.
Proof.
  intros x H y.
  setoid_rewrite H at 1 2.
  setoid_rewrite <- (H y) at 1 2.
  setoid_rewrite H.
  reflexivity.
Qed.

Axiom x : unit.
Axiom H : forall y : unit, y = x.
Goal forall y : unit, y = y.
Proof.
  assert (H2 : tt = x).
  { setoid_rewrite H with (y := tt).
    reflexivity. }
  setoid_rewrite H with (y := tt) at 1 in H2.
  intros y.
  setoid_rewrite H with (y := y) at 1.
  setoid_rewrite H with (y := y) at 1.
  reflexivity.
Qed.


(* view intro patterns *)
Axiom lem : forall x:nat, bool -> x = x.

Axiom lem' : 3 = 3 -> False.

Goal bool -> 2 = 2.
  Fail intros a%lem.
  eintros a%lem.
  revert a.
  intros b%lem'.
  destruct b.
Qed.

(* assert interns and typechecks the type as a type
   (bug #15094) *)

Inductive my_type :=
| my_el : my_type.

Definition my_coercion (_ : my_type) := True.

Coercion my_coercion : my_type >-> Sortclass.

Goal False.
  (* expected type is sortclass -> coercion inserted *)
  assert my_el by exact I.

  (* interned as with type scope locally open *)
  assert (H' : nat * (0 * 0 = 0)) by (split; first [exact 0 | reflexivity]).
Abort.

Goal nat.
  Std.apply true false [fun () => Control.plus (fun () => 'I) (fun _ => '0), Std.NoBindings] None.
Qed.