-
Notifications
You must be signed in to change notification settings - Fork 83
/
LogicProgramming_template.v
282 lines (202 loc) · 5.17 KB
/
LogicProgramming_template.v
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
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
(** Formal Reasoning About Programs <http://adam.chlipala.net/frap/>
* Supplementary Coq material: unification and logic programming
* Author: Adam Chlipala
* License: https://creativecommons.org/licenses/by-nc-nd/4.0/
* Much of the material comes from CPDT <http://adam.chlipala.net/cpdt/> by the same author. *)
Require Import Frap.
Set Implicit Arguments.
(** * Introducing Logic Programming *)
(* Recall the definition of addition from the standard library. *)
Definition real_plus := Eval compute in plus.
Print real_plus.
(* Alternatively, we can define it as a relation. *)
Inductive plusR : nat -> nat -> nat -> Prop :=
| PlusO : forall m, plusR O m m
| PlusS : forall n m r, plusR n m r
-> plusR (S n) m (S r).
(* Let's prove the correspondence. *)
Theorem plusR_plus : forall n m r,
plusR n m r
-> r = n + m.
Proof.
Admitted.
Theorem plus_plusR : forall n m,
plusR n m (n + m).
Proof.
Admitted.
Example four_plus_three : 4 + 3 = 7.
Proof.
reflexivity.
Qed.
Print four_plus_three.
Example four_plus_three' : plusR 4 3 7.
Proof.
Admitted.
Print four_plus_three'.
Example five_plus_three : plusR 5 3 8.
Proof.
Admitted.
(* Demonstrating _backtracking_ *)
Example seven_minus_three : exists x, x + 3 = 7.
Proof.
apply ex_intro with 0.
Abort.
Example seven_minus_three' : exists x, plusR x 3 7.
Proof.
Admitted.
(* Backwards! *)
Example seven_minus_four' : exists x, plusR 4 x 7.
Proof.
Admitted.
Example seven_minus_three'' : exists x, x + 3 = 7.
Proof.
Admitted.
Example seven_minus_four : exists x, 4 + x = 7.
Proof.
Admitted.
Example seven_minus_four_zero : exists x, 4 + x + 0 = 7.
Proof.
Admitted.
Check eq_trans.
Section slow.
Hint Resolve eq_trans : core.
Example zero_minus_one : exists x, 1 + x = 0.
Time eauto 1.
Time eauto 2.
Time eauto 3.
Time eauto 4.
Time eauto 5.
debug eauto 3.
Abort.
End slow.
Example from_one_to_zero : exists x, 1 + x = 0.
Proof.
Admitted.
Example seven_minus_three_again : exists x, x + 3 = 7.
Proof.
Admitted.
Example needs_trans : forall x y, 1 + x = y
-> y = 2
-> exists z, z + x = 3.
Proof.
Admitted.
(** * Searching for Underconstrained Values *)
Print Datatypes.length.
Example length_1_2 : length (1 :: 2 :: nil) = 2.
Proof.
Admitted.
Print length_1_2.
Example length_is_2 : exists ls : list nat, length ls = 2.
Proof.
Abort.
Print Forall.
Example length_is_2 : exists ls : list nat, length ls = 2
/\ Forall (fun n => n >= 1) ls.
Proof.
Admitted.
Print length_is_2.
Definition sum := fold_right plus O.
Example length_and_sum : exists ls : list nat, length ls = 2
/\ sum ls = O.
Proof.
Admitted.
Print length_and_sum.
Example length_and_sum' : exists ls : list nat, length ls = 5
/\ sum ls = 42.
Proof.
Admitted.
Print length_and_sum'.
Example length_and_sum'' : exists ls : list nat, length ls = 2
/\ sum ls = 3
/\ Forall (fun n => n <> 0) ls.
Proof.
Admitted.
Print length_and_sum''.
(** * Synthesizing Programs *)
Inductive exp : Set :=
| Const (n : nat)
| Var
| Plus (e1 e2 : exp).
Inductive eval (var : nat) : exp -> nat -> Prop :=
| EvalConst : forall n, eval var (Const n) n
| EvalVar : eval var Var var
| EvalPlus : forall e1 e2 n1 n2, eval var e1 n1
-> eval var e2 n2
-> eval var (Plus e1 e2) (n1 + n2).
Local Hint Constructors eval : core.
Example eval1 : forall var, eval var (Plus Var (Plus (Const 8) Var)) (var + (8 + var)).
Proof.
auto.
Qed.
Example eval1' : forall var, eval var (Plus Var (Plus (Const 8) Var)) (2 * var + 8).
Proof.
eauto.
Abort.
Example eval1' : forall var, eval var (Plus Var (Plus (Const 8) Var)) (2 * var + 8).
Proof.
Admitted.
Example synthesize1 : exists e, forall var, eval var e (var + 7).
Proof.
Admitted.
Print synthesize1.
(* Here are two more examples showing off our program-synthesis abilities. *)
Example synthesize2 : exists e, forall var, eval var e (2 * var + 8).
Proof.
Admitted.
Print synthesize2.
Example synthesize3 : exists e, forall var, eval var e (3 * var + 42).
Proof.
Admitted.
Print synthesize3.
Theorem linear : forall e, exists k n,
forall var, eval var e (k * var + n).
Proof.
Admitted.
Section side_effect_sideshow.
Variable A : Set.
Variables P Q : A -> Prop.
Variable x : A.
Hypothesis Px : P x.
Hypothesis Qx : Q x.
Theorem double_threat : exists y, P y /\ Q y.
Proof.
eexists; propositional.
eauto.
eauto.
Qed.
End side_effect_sideshow.
(** * More on [auto] Hints *)
Theorem bool_neq : true <> false.
Proof.
Admitted.
Section forall_and.
Variable A : Set.
Variables P Q : A -> Prop.
Hypothesis both : forall x, P x /\ Q x.
Theorem forall_and : forall z, P z.
Proof.
Admitted.
End forall_and.
(** * Rewrite Hints *)
Section autorewrite.
Variable A : Set.
Variable f : A -> A.
Hypothesis f_f : forall x, f (f x) = f x.
Hint Rewrite f_f.
Lemma f_f_f : forall x, f (f (f x)) = f x.
Proof.
intros; autorewrite with core; reflexivity.
Qed.
Section garden_path.
Variable g : A -> A.
Hypothesis f_g : forall x, f x = g x.
Hint Rewrite f_g.
Lemma f_f_f' : forall x, f (f (f x)) = f x.
Proof.
Admitted.
End garden_path.
Lemma in_star : forall x y, f (f (f (f x))) = f (f y)
-> f x = f (f (f y)).
Proof.
Admitted.
End autorewrite.