comparison src/Reflection.v @ 221:54e8ecb5ec88

Port Reflection
author Adam Chlipala <adamc@hcoop.net>
date Mon, 16 Nov 2009 11:09:47 -0500
parents f05514cc6c0d
children 4662b6f099b0
comparison
equal deleted inserted replaced
220:15501145d696 221:54e8ecb5ec88
1 (* Copyright (c) 2008, Adam Chlipala 1 (* Copyright (c) 2008-2009, Adam Chlipala
2 * 2 *
3 * This work is licensed under a 3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License. 5 * Unported License.
6 * The license text is available at: 6 * The license text is available at:
36 Theorem even_256 : isEven 256. 36 Theorem even_256 : isEven 256.
37 prove_even. 37 prove_even.
38 Qed. 38 Qed.
39 39
40 Print even_256. 40 Print even_256.
41 (** [[ 41 (** %\vspace{-.15in}% [[
42
43 even_256 = 42 even_256 =
44 Even_SS 43 Even_SS
45 (Even_SS 44 (Even_SS
46 (Even_SS 45 (Even_SS
47 (Even_SS 46 (Even_SS
48 ]] 47
49 48 ]]
50 ...and so on. This procedure always works (at least on machines with infinite resources), but it has a serious drawback, which we see when we print the proof it generates that 256 is even. The final proof term has length linear in the input value. This seems like a shame, since we could write a trivial and trustworthy program to verify evenness of constants. The proof checker could simply call our program where needed. 49
50 %\noindent%...and so on. This procedure always works (at least on machines with infinite resources), but it has a serious drawback, which we see when we print the proof it generates that 256 is even. The final proof term has length linear in the input value. This seems like a shame, since we could write a trivial and trustworthy program to verify evenness of constants. The proof checker could simply call our program where needed.
51 51
52 It is also unfortunate not to have static typing guarantees that our tactic always behaves appropriately. Other invocations of similar tactics might fail with dynamic type errors, and we would not know about the bugs behind these errors until we happened to attempt to prove complex enough goals. 52 It is also unfortunate not to have static typing guarantees that our tactic always behaves appropriately. Other invocations of similar tactics might fail with dynamic type errors, and we would not know about the bugs behind these errors until we happened to attempt to prove complex enough goals.
53 53
54 The techniques of proof by reflection address both complaints. We will be able to write proofs like this with constant size overhead beyond the size of the input, and we will do it with verified decision procedures written in Gallina. 54 The techniques of proof by reflection address both complaints. We will be able to write proofs like this with constant size overhead beyond the size of the input, and we will do it with verified decision procedures written in Gallina.
55 55
56 For this example, we begin by using a type from the [MoreSpecif] module to write a certified evenness checker. *) 56 For this example, we begin by using a type from the [MoreSpecif] module (included in the book source) to write a certified evenness checker. *)
57 57
58 Print partial. 58 Print partial.
59 (** [[ 59 (** %\vspace{-.15in}% [[
60
61 Inductive partial (P : Prop) : Set := Proved : P -> [P] | Uncertain : [P] 60 Inductive partial (P : Prop) : Set := Proved : P -> [P] | Uncertain : [P]
62 ]] *) 61
63 62 ]]
64 (** A [partial P] value is an optional proof of [P]. The notation [[P]] stands for [partial P]. *) 63
65 64 A [partial P] value is an optional proof of [P]. The notation [[P]] stands for [partial P]. *)
66 Open Local Scope partial_scope. 65
66 Local Open Scope partial_scope.
67 67
68 (** We bring into scope some notations for the [partial] type. These overlap with some of the notations we have seen previously for specification types, so they were placed in a separate scope that needs separate opening. *) 68 (** We bring into scope some notations for the [partial] type. These overlap with some of the notations we have seen previously for specification types, so they were placed in a separate scope that needs separate opening. *)
69 69
70 (* begin thide *) 70 (* begin thide *)
71 Definition check_even (n : nat) : [isEven n]. 71 Definition check_even (n : nat) : [isEven n].
72 Hint Constructors isEven. 72 Hint Constructors isEven.
73 73
74 refine (fix F (n : nat) : [isEven n] := 74 refine (fix F (n : nat) : [isEven n] :=
75 match n return [isEven n] with 75 match n with
76 | 0 => Yes 76 | 0 => Yes
77 | 1 => No 77 | 1 => No
78 | S (S n') => Reduce (F n') 78 | S (S n') => Reduce (F n')
79 end); auto. 79 end); auto.
80 Defined. 80 Defined.
103 Theorem even_256' : isEven 256. 103 Theorem even_256' : isEven 256.
104 prove_even_reflective. 104 prove_even_reflective.
105 Qed. 105 Qed.
106 106
107 Print even_256'. 107 Print even_256'.
108 (** [[ 108 (** %\vspace{-.15in}% [[
109
110 even_256' = partialOut (check_even 256) 109 even_256' = partialOut (check_even 256)
111 : isEven 256 110 : isEven 256
111
112 ]] 112 ]]
113 113
114 We can see a constant wrapper around the object of the proof. For any even number, this form of proof will suffice. What happens if we try the tactic with an odd number? *) 114 We can see a constant wrapper around the object of the proof. For any even number, this form of proof will suffice. What happens if we try the tactic with an odd number? *)
115 115
116 Theorem even_255 : isEven 255. 116 Theorem even_255 : isEven 255.
117 (** [[ 117 (** [[
118
119 prove_even_reflective. 118 prove_even_reflective.
120 119
120 User error: No matching clauses for match goal
121
121 ]] 122 ]]
122 123
124 Thankfully, the tactic fails. To see more precisely what goes wrong, we can run manually the body of the [match].
125
123 [[ 126 [[
124
125 User error: No matching clauses for match goal
126 ]]
127
128 Thankfully, the tactic fails. To see more precisely what goes wrong, we can run manually the body of the [match].
129
130 [[
131
132 exact (partialOut (check_even 255)). 127 exact (partialOut (check_even 255)).
133
134 ]]
135
136 [[
137 128
138 Error: The term "partialOut (check_even 255)" has type 129 Error: The term "partialOut (check_even 255)" has type
139 "match check_even 255 with 130 "match check_even 255 with
140 | Yes => isEven 255 131 | Yes => isEven 255
141 | No => True 132 | No => True
142 end" while it is expected to have type "isEven 255" 133 end" while it is expected to have type "isEven 255"
134
143 ]] 135 ]]
144 136
145 As usual, the type-checker performs no reductions to simplify error messages. If we reduced the first term ourselves, we would see that [check_even 255] reduces to a [No], so that the first term is equivalent to [True], which certainly does not unify with [isEven 255]. *) 137 As usual, the type-checker performs no reductions to simplify error messages. If we reduced the first term ourselves, we would see that [check_even 255] reduces to a [No], so that the first term is equivalent to [True], which certainly does not unify with [isEven 255]. *)
138
146 Abort. 139 Abort.
147 140
148 141
149 (** * Reflecting the Syntax of a Trivial Tautology Language *) 142 (** * Reflecting the Syntax of a Trivial Tautology Language *)
150 143
153 Theorem true_galore : (True /\ True) -> (True \/ (True /\ (True -> True))). 146 Theorem true_galore : (True /\ True) -> (True \/ (True /\ (True -> True))).
154 tauto. 147 tauto.
155 Qed. 148 Qed.
156 149
157 Print true_galore. 150 Print true_galore.
158 151 (** %\vspace{-.15in}% [[
159 (** [[
160
161 true_galore = 152 true_galore =
162 fun H : True /\ True => 153 fun H : True /\ True =>
163 and_ind (fun _ _ : True => or_introl (True /\ (True -> True)) I) H 154 and_ind (fun _ _ : True => or_introl (True /\ (True -> True)) I) H
164 : True /\ True -> True \/ True /\ (True -> True) 155 : True /\ True -> True \/ True /\ (True -> True)
156
165 ]] 157 ]]
166 158
167 As we might expect, the proof that [tauto] builds contains explicit applications of natural deduction rules. For large formulas, this can add a linear amount of proof size overhead, beyond the size of the input. 159 As we might expect, the proof that [tauto] builds contains explicit applications of natural deduction rules. For large formulas, this can add a linear amount of proof size overhead, beyond the size of the input.
168 160
169 To write a reflective procedure for this class of goals, we will need to get into the actual "reflection" part of "proof by reflection." It is impossible to case-analyze a [Prop] in any way in Gallina. We must %\textit{%#<i>#reflect#</i>#%}% [Prop] into some type that we %\textit{%#<i>#can#</i>#%}% analyze. This inductive type is a good candidate: *) 161 To write a reflective procedure for this class of goals, we will need to get into the actual "reflection" part of "proof by reflection." It is impossible to case-analyze a [Prop] in any way in Gallina. We must %\textit{%#<i>#reflect#</i>#%}% [Prop] into some type that we %\textit{%#<i>#can#</i>#%}% analyze. This inductive type is a good candidate: *)
226 obvious. 218 obvious.
227 Qed. 219 Qed.
228 220
229 Print true_galore'. 221 Print true_galore'.
230 222
231 (** [[ 223 (** %\vspace{-.15in}% [[
232
233 true_galore' = 224 true_galore' =
234 tautTrue 225 tautTrue
235 (TautImp (TautAnd TautTrue TautTrue) 226 (TautImp (TautAnd TautTrue TautTrue)
236 (TautOr TautTrue (TautAnd TautTrue (TautImp TautTrue TautTrue)))) 227 (TautOr TautTrue (TautAnd TautTrue (TautImp TautTrue TautTrue))))
237 : True /\ True -> True \/ True /\ (True -> True) 228 : True /\ True -> True \/ True /\ (True -> True)
238 229
239 ]] 230 ]]
240 231
241 It is worth considering how the reflective tactic improves on a pure-Ltac implementation. The formula reflection process is just as ad-hoc as before, so we gain little there. In general, proofs will be more complicated than formula translation, and the "generic proof rule" that we apply here %\textit{%#<i>#is#</i>#%}% on much better formal footing than a recursive Ltac function. The dependent type of the proof guarantees that it "works" on any input formula. This is all in addition to the proof-size improvement that we have already seen. *) 232 It is worth considering how the reflective tactic improves on a pure-Ltac implementation. The formula reflection process is just as ad-hoc as before, so we gain little there. In general, proofs will be more complicated than formula translation, and the "generic proof rule" that we apply here %\textit{%#<i>#is#</i>#%}% on much better formal footing than a recursive Ltac function. The dependent type of the proof guarantees that it "works" on any input formula. This is all in addition to the proof-size improvement that we have already seen. *)
242 233
243 234
341 (* end thide *) 332 (* end thide *)
342 333
343 Theorem t1 : forall a b c d, a + b + c + d = a + (b + c) + d. 334 Theorem t1 : forall a b c d, a + b + c + d = a + (b + c) + d.
344 intros; monoid. 335 intros; monoid.
345 (** [[ 336 (** [[
346
347 ============================ 337 ============================
348 a + (b + (c + (d + e))) = a + (b + (c + (d + e))) 338 a + (b + (c + (d + e))) = a + (b + (c + (d + e)))
339
349 ]] 340 ]]
350 341
351 [monoid] has canonicalized both sides of the equality, such that we can finish the proof by reflexivity. *) 342 [monoid] has canonicalized both sides of the equality, such that we can finish the proof by reflexivity. *)
352 343
353 reflexivity. 344 reflexivity.
354 Qed. 345 Qed.
355 346
356 (** It is interesting to look at the form of the proof. *) 347 (** It is interesting to look at the form of the proof. *)
357 348
358 Print t1. 349 Print t1.
359 (** [[ 350 (** %\vspace{-.15in}% [[
360
361 t1 = 351 t1 =
362 fun a b c d : A => 352 fun a b c d : A =>
363 monoid_reflect (Op (Op (Op (Var a) (Var b)) (Var c)) (Var d)) 353 monoid_reflect (Op (Op (Op (Var a) (Var b)) (Var c)) (Var d))
364 (Op (Op (Var a) (Op (Var b) (Var c))) (Var d)) 354 (Op (Op (Var a) (Op (Var b) (Var c))) (Var d))
365 (refl_equal (a + (b + (c + (d + e))))) 355 (refl_equal (a + (b + (c + (d + e)))))
366 : forall a b c d : A, a + b + c + d = a + (b + c) + d 356 : forall a b c d : A, a + b + c + d = a + (b + c) + d
357
367 ]] 358 ]]
368 359
369 The proof term contains only restatements of the equality operands in reflected form, followed by a use of reflexivity on the shared canonical form. *) 360 The proof term contains only restatements of the equality operands in reflected form, followed by a use of reflexivity on the shared canonical form. *)
361
370 End monoid. 362 End monoid.
371 363
372 (** Extensions of this basic approach are used in the implementations of the [ring] and [field] tactics that come packaged with Coq. *) 364 (** Extensions of this basic approach are used in the implementations of the [ring] and [field] tactics that come packaged with Coq. *)
373 365
374 366
375 (** * A Smarter Tautology Solver *) 367 (** * A Smarter Tautology Solver *)
376 368
377 (** Now we are ready to revisit our earlier tautology solver example. We want to broaden the scope of the tactic to include formulas whose truth is not syntactically apparent. We will want to allow injection of arbitrary formulas, like we allowed arbitrary monoid expressions in the last example. Since we are working in a richer theory, it is important to be able to use equalities between different injected formulas. For instance, we cannott prove [P -> P] by translating the formula into a value like [Imp (Var P) (Var P)], because a Gallina function has no way of comparing the two [P]s for equality. 369 (** Now we are ready to revisit our earlier tautology solver example. We want to broaden the scope of the tactic to include formulas whose truth is not syntactically apparent. We will want to allow injection of arbitrary formulas, like we allowed arbitrary monoid expressions in the last example. Since we are working in a richer theory, it is important to be able to use equalities between different injected formulas. For instance, we cannot prove [P -> P] by translating the formula into a value like [Imp (Var P) (Var P)], because a Gallina function has no way of comparing the two [P]s for equality.
378 370
379 To arrive at a nice implementation satisfying these criteria, we introduce the [quote] tactic and its associated library. *) 371 To arrive at a nice implementation satisfying these criteria, we introduce the [quote] tactic and its associated library. *)
380 372
381 Require Import Quote. 373 Require Import Quote.
382 374
425 decide equality. 417 decide equality.
426 Defined. 418 Defined.
427 419
428 Definition add (s : set index) (v : index) := set_add index_eq v s. 420 Definition add (s : set index) (v : index) := set_add index_eq v s.
429 421
430 Definition In_dec : forall v (s : set index), {In v s} + {~In v s}. 422 Definition In_dec : forall v (s : set index), {In v s} + {~ In v s}.
431 Open Local Scope specif_scope. 423 Local Open Scope specif_scope.
432 424
433 intro; refine (fix F (s : set index) : {In v s} + {~In v s} := 425 intro; refine (fix F (s : set index) : {In v s} + {~ In v s} :=
434 match s return {In v s} + {~In v s} with 426 match s with
435 | nil => No 427 | nil => No
436 | v' :: s' => index_eq v' v || F s' 428 | v' :: s' => index_eq v' v || F s'
437 end); crush. 429 end); crush.
438 Defined. 430 Defined.
439 431
462 induction s; crush. 454 induction s; crush.
463 Qed. 455 Qed.
464 456
465 Hint Resolve allTrue_add allTrue_In. 457 Hint Resolve allTrue_add allTrue_In.
466 458
467 Open Local Scope partial_scope. 459 Local Open Scope partial_scope.
468 460
469 (** Now we can write a function [forward] which implements deconstruction of hypotheses. It has a dependent type, in the style of Chapter 6, guaranteeing correctness. The arguments to [forward] are a goal formula [f], a set [known] of atomic formulas that we may assume are true, a hypothesis formula [hyp], and a success continuation [cont] that we call when we have extended [known] to hold new truths implied by [hyp]. *) 461 (** Now we can write a function [forward] which implements deconstruction of hypotheses. It has a dependent type, in the style of Chapter 6, guaranteeing correctness. The arguments to [forward] are a goal formula [f], a set [known] of atomic formulas that we may assume are true, a hypothesis formula [hyp], and a success continuation [cont] that we call when we have extended [known] to hold new truths implied by [hyp]. *)
470 462
471 Definition forward (f : formula) (known : set index) (hyp : formula) 463 Definition forward (f : formula) (known : set index) (hyp : formula)
472 (cont : forall known', [allTrue known' -> formulaDenote atomics f]) 464 (cont : forall known', [allTrue known' -> formulaDenote atomics f])
473 : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f]. 465 : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f].
474 refine (fix F (f : formula) (known : set index) (hyp : formula) 466 refine (fix F (f : formula) (known : set index) (hyp : formula)
475 (cont : forall known', [allTrue known' -> formulaDenote atomics f]){struct hyp} 467 (cont : forall known', [allTrue known' -> formulaDenote atomics f])
476 : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f] := 468 : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f] :=
477 match hyp return [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f] with 469 match hyp with
478 | Atomic v => Reduce (cont (add known v)) 470 | Atomic v => Reduce (cont (add known v))
479 | Truth => Reduce (cont known) 471 | Truth => Reduce (cont known)
480 | Falsehood => Yes 472 | Falsehood => Yes
481 | And h1 h2 => 473 | And h1 h2 =>
482 Reduce (F (Imp h2 f) known h1 (fun known' => 474 Reduce (F (Imp h2 f) known h1 (fun known' =>
486 end); crush. 478 end); crush.
487 Defined. 479 Defined.
488 480
489 (** A [backward] function implements analysis of the final goal. It calls [forward] to handle implications. *) 481 (** A [backward] function implements analysis of the final goal. It calls [forward] to handle implications. *)
490 482
491 Definition backward (known : set index) (f : formula) : [allTrue known -> formulaDenote atomics f]. 483 Definition backward (known : set index) (f : formula)
492 refine (fix F (known : set index) (f : formula) : [allTrue known -> formulaDenote atomics f] := 484 : [allTrue known -> formulaDenote atomics f].
493 match f return [allTrue known -> formulaDenote atomics f] with 485 refine (fix F (known : set index) (f : formula)
486 : [allTrue known -> formulaDenote atomics f] :=
487 match f with
494 | Atomic v => Reduce (In_dec v known) 488 | Atomic v => Reduce (In_dec v known)
495 | Truth => Yes 489 | Truth => Yes
496 | Falsehood => No 490 | Falsehood => No
497 | And f1 f2 => F known f1 && F known f2 491 | And f1 f2 => F known f1 && F known f2
498 | Or f1 f2 => F known f1 || F known f2 492 | Or f1 f2 => F known f1 || F known f2
528 Theorem mt1 : True. 522 Theorem mt1 : True.
529 my_tauto. 523 my_tauto.
530 Qed. 524 Qed.
531 525
532 Print mt1. 526 Print mt1.
533 (** [[ 527 (** %\vspace{-.15in}% [[
534
535 mt1 = partialOut (my_tauto (Empty_vm Prop) Truth) 528 mt1 = partialOut (my_tauto (Empty_vm Prop) Truth)
536 : True 529 : True
530
537 ]] 531 ]]
538 532
539 We see [my_tauto] applied with an empty [varmap], since every subformula is handled by [formulaDenote]. *) 533 We see [my_tauto] applied with an empty [varmap], since every subformula is handled by [formulaDenote]. *)
540 534
541 Theorem mt2 : forall x y : nat, x = y --> x = y. 535 Theorem mt2 : forall x y : nat, x = y --> x = y.
542 my_tauto. 536 my_tauto.
543 Qed. 537 Qed.
544 538
545 Print mt2. 539 Print mt2.
546 (** [[ 540 (** %\vspace{-.15in}% [[
547
548 mt2 = 541 mt2 =
549 fun x y : nat => 542 fun x y : nat =>
550 partialOut 543 partialOut
551 (my_tauto (Node_vm (x = y) (Empty_vm Prop) (Empty_vm Prop)) 544 (my_tauto (Node_vm (x = y) (Empty_vm Prop) (Empty_vm Prop))
552 (Imp (Atomic End_idx) (Atomic End_idx))) 545 (Imp (Atomic End_idx) (Atomic End_idx)))
553 : forall x y : nat, x = y --> x = y 546 : forall x y : nat, x = y --> x = y
547
554 ]] 548 ]]
555 549
556 Crucially, both instances of [x = y] are represented with the same index, [End_idx]. The value of this index only needs to appear once in the [varmap], whose form reveals that [varmap]s are represented as binary trees, where [index] values denote paths from tree roots to leaves. *) 550 Crucially, both instances of [x = y] are represented with the same index, [End_idx]. The value of this index only needs to appear once in the [varmap], whose form reveals that [varmap]s are represented as binary trees, where [index] values denote paths from tree roots to leaves. *)
557 551
558 Theorem mt3 : forall x y z, 552 Theorem mt3 : forall x y z,
560 --> y > z /\ (x < y \/ x < S y). 554 --> y > z /\ (x < y \/ x < S y).
561 my_tauto. 555 my_tauto.
562 Qed. 556 Qed.
563 557
564 Print mt3. 558 Print mt3.
565 (** [[ 559 (** %\vspace{-.15in}% [[
566
567 fun x y z : nat => 560 fun x y z : nat =>
568 partialOut 561 partialOut
569 (my_tauto 562 (my_tauto
570 (Node_vm (x < S y) (Node_vm (x < y) (Empty_vm Prop) (Empty_vm Prop)) 563 (Node_vm (x < S y) (Node_vm (x < y) (Empty_vm Prop) (Empty_vm Prop))
571 (Node_vm (y > z) (Empty_vm Prop) (Empty_vm Prop))) 564 (Node_vm (y > z) (Empty_vm Prop) (Empty_vm Prop)))
574 (And (Atomic (Right_idx End_idx)) (Atomic End_idx))) 567 (And (Atomic (Right_idx End_idx)) (Atomic End_idx)))
575 (And (Atomic (Right_idx End_idx)) 568 (And (Atomic (Right_idx End_idx))
576 (Or (Atomic (Left_idx End_idx)) (Atomic End_idx))))) 569 (Or (Atomic (Left_idx End_idx)) (Atomic End_idx)))))
577 : forall x y z : nat, 570 : forall x y z : nat,
578 x < y /\ y > z \/ y > z /\ x < S y --> y > z /\ (x < y \/ x < S y) 571 x < y /\ y > z \/ y > z /\ x < S y --> y > z /\ (x < y \/ x < S y)
572
579 ]] 573 ]]
580 574
581 Our goal contained three distinct atomic formulas, and we see that a three-element [varmap] is generated. 575 Our goal contained three distinct atomic formulas, and we see that a three-element [varmap] is generated.
582 576
583 It can be interesting to observe differences between the level of repetition in proof terms generated by [my_tauto] and [tauto] for especially trivial theorems. *) 577 It can be interesting to observe differences between the level of repetition in proof terms generated by [my_tauto] and [tauto] for especially trivial theorems. *)
585 Theorem mt4 : True /\ True /\ True /\ True /\ True /\ True /\ False --> False. 579 Theorem mt4 : True /\ True /\ True /\ True /\ True /\ True /\ False --> False.
586 my_tauto. 580 my_tauto.
587 Qed. 581 Qed.
588 582
589 Print mt4. 583 Print mt4.
590 (** [[ 584 (** %\vspace{-.15in}% [[
591
592 mt4 = 585 mt4 =
593 partialOut 586 partialOut
594 (my_tauto (Empty_vm Prop) 587 (my_tauto (Empty_vm Prop)
595 (Imp 588 (Imp
596 (And Truth 589 (And Truth
603 Theorem mt4' : True /\ True /\ True /\ True /\ True /\ True /\ False -> False. 596 Theorem mt4' : True /\ True /\ True /\ True /\ True /\ True /\ False -> False.
604 tauto. 597 tauto.
605 Qed. 598 Qed.
606 599
607 Print mt4'. 600 Print mt4'.
608 (** [[ 601 (** %\vspace{-.15in}% [[
609
610 mt4' = 602 mt4' =
611 fun H : True /\ True /\ True /\ True /\ True /\ True /\ False => 603 fun H : True /\ True /\ True /\ True /\ True /\ True /\ False =>
612 and_ind 604 and_ind
613 (fun (_ : True) (H1 : True /\ True /\ True /\ True /\ True /\ False) => 605 (fun (_ : True) (H1 : True /\ True /\ True /\ True /\ True /\ False) =>
614 and_ind 606 and_ind
625 ]] *) 617 ]] *)
626 618
627 619
628 (** * Exercises *) 620 (** * Exercises *)
629 621
622 (** remove printing * *)
623
630 (** %\begin{enumerate}%#<ol># 624 (** %\begin{enumerate}%#<ol>#
631 625
632 %\item%#<li># Implement a reflective procedure for normalizing systems of linear equations over rational numbers. In particular, the tactic should identify all hypotheses that are linear equations over rationals where the equation righthand sides are constants. It should normalize each hypothesis to have a lefthand side that is a sum of products of constants and variables, with no variable appearing multiple times. Then, your tactic should add together all of these equations to form a single new equation, possibly clearing the original equations. Some coefficients may cancel in the addition, reducing the number of variables that appear. 626 %\item%#<li># Implement a reflective procedure for normalizing systems of linear equations over rational numbers. In particular, the tactic should identify all hypotheses that are linear equations over rationals where the equation righthand sides are constants. It should normalize each hypothesis to have a lefthand side that is a sum of products of constants and variables, with no variable appearing multiple times. Then, your tactic should add together all of these equations to form a single new equation, possibly clearing the original equations. Some coefficients may cancel in the addition, reducing the number of variables that appear.
633 627
634 To work with rational numbers, import module [QArith] and use [Open Local Scope Q_scope]. All of the usual arithmetic operator notations will then work with rationals, and there are shorthands for constants 0 and 1. Other rationals must be written as [num # den] for numerator [num] and denominator [den]. Use the infix operator [==] in place of [=], to deal with different ways of expressing the same number as a fraction. For instance, a theorem and proof like this one should work with your tactic: 628 To work with rational numbers, import module [QArith] and use [Local Open Scope Q_scope]. All of the usual arithmetic operator notations will then work with rationals, and there are shorthands for constants 0 and 1. Other rationals must be written as [num # den] for numerator [num] and denominator [den]. Use the infix operator [==] in place of [=], to deal with different ways of expressing the same number as a fraction. For instance, a theorem and proof like this one should work with your tactic:
635 629
636 [[ 630 [[
637 Theorem t2 : forall x y z, (2 # 1) * (x - (3 # 2) * y) == 15 # 1 631 Theorem t2 : forall x y z, (2 # 1) * (x - (3 # 2) * y) == 15 # 1
638 -> z + (8 # 1) * x == 20 # 1 632 -> z + (8 # 1) * x == 20 # 1
639 -> (-6 # 2) * y + (10 # 1) * x + z == 35 # 1. 633 -> (-6 # 2) * y + (10 # 1) * x + z == 35 # 1.
640
641 ]]
642
643 [[
644
645 intros; reflectContext; assumption. 634 intros; reflectContext; assumption.
646
647 ]]
648 [[
649 Qed. 635 Qed.
650 636
651 ]] 637 ]]
652 638
653 Your solution can work in any way that involves reflecting syntax and doing most calculation with a Gallina function. These hints outline a particular possible solution. Throughout, the [ring] tactic will be helpful for proving many simple facts about rationals, and tactics like [rewrite] are correctly overloaded to work with rational equality [==]. 639 Your solution can work in any way that involves reflecting syntax and doing most calculation with a Gallina function. These hints outline a particular possible solution. Throughout, the [ring] tactic will be helpful for proving many simple facts about rationals, and tactics like [rewrite] are correctly overloaded to work with rational equality [==].
654 640
655 %\begin{enumerate}%#<ol># 641 %\begin{enumerate}%#<ol>#
656 %\item%#<li># Define an inductive type [exp] of expressions over rationals (which inhabit the Coq type [Q]). Include variables (represented as natural numbers), constants, addition, substraction, and multiplication.#</li># 642 %\item%#<li># Define an inductive type [exp] of expressions over rationals (which inhabit the Coq type [Q]). Include variables (represented as natural numbers), constants, addition, subtraction, and multiplication.#</li>#
657 %\item%#<li># Define a function [lookup] for reading an element out of a list of rationals, by its position in the list.#</li># 643 %\item%#<li># Define a function [lookup] for reading an element out of a list of rationals, by its position in the list.#</li>#
658 %\item%#<li># Define a function [expDenote] that translates [exp]s, along with lists of rationals representing variable values, to [Q].#</li># 644 %\item%#<li># Define a function [expDenote] that translates [exp]s, along with lists of rationals representing variable values, to [Q].#</li>#
659 %\item%#<li># Define a recursive function [eqsDenote] over [list (exp * Q)], characterizing when all of the equations are true.#</li># 645 %\item%#<li># Define a recursive function [eqsDenote] over [list (exp * Q)], characterizing when all of the equations are true.#</li>#
660 %\item%#<li># Fix a representation [lhs] of flattened expressions. Where [len] is the number of variables, represent a flattened equation as [ilist Q len]. Each position of the list gives the coefficient of the corresponding variable.#</li># 646 %\item%#<li># Fix a representation [lhs] of flattened expressions. Where [len] is the number of variables, represent a flattened equation as [ilist Q len]. Each position of the list gives the coefficient of the corresponding variable.#</li>#
661 %\item%#<li># Write a recursive function [linearize] that takes a constant [k] and an expression [e] and optionally returns an [lhs] equivalent to [k * e]. This function returns [None] when it discovers that the input expression is not linear. The parameter [len] of [lhs] should be a parameter of [linearize], too. The functions [singleton], [everywhere], and [map2] from [DepList] will probably be helpful. It is also helpful to know that [Qplus] is the identifier for rational addition.#</li># 647 %\item%#<li># Write a recursive function [linearize] that takes a constant [k] and an expression [e] and optionally returns an [lhs] equivalent to [k * e]. This function returns [None] when it discovers that the input expression is not linear. The parameter [len] of [lhs] should be a parameter of [linearize], too. The functions [singleton], [everywhere], and [map2] from [DepList] will probably be helpful. It is also helpful to know that [Qplus] is the identifier for rational addition.#</li>#
665 %\item%#<li># Prove that, when [linearizeEqs] succeeds on an equation list [eqs], then the final summed-up equation is true whenever the original equation list is true.#</li># 651 %\item%#<li># Prove that, when [linearizeEqs] succeeds on an equation list [eqs], then the final summed-up equation is true whenever the original equation list is true.#</li>#
666 %\item%#<li># Write a tactic [findVarsHyps] to search through all equalities on rationals in the context, recursing through addition, subtraction, and multiplication to find the list of expressions that should be treated as variables. This list should be suitable as an argument to [expDenote] and [eqsDenote], associating a [Q] value to each natural number that stands for a variable.#</li># 652 %\item%#<li># Write a tactic [findVarsHyps] to search through all equalities on rationals in the context, recursing through addition, subtraction, and multiplication to find the list of expressions that should be treated as variables. This list should be suitable as an argument to [expDenote] and [eqsDenote], associating a [Q] value to each natural number that stands for a variable.#</li>#
667 %\item%#<li># Write a tactic [reflect] to reflect a [Q] expression into [exp], with respect to a given list of variable values.#</li># 653 %\item%#<li># Write a tactic [reflect] to reflect a [Q] expression into [exp], with respect to a given list of variable values.#</li>#
668 %\item%#<li># Write a tactic [reflectEqs] to reflect a formula that begins with a sequence of implications from linear equalities whose lefthand sides are expressed with [expDenote]. This tactic should build a [list (exp * Q)] representing the equations. Remember to give an explicit type annotation when returning a nil list, as in [constr:(@nil (exp * Q))].#</li># 654 %\item%#<li># Write a tactic [reflectEqs] to reflect a formula that begins with a sequence of implications from linear equalities whose lefthand sides are expressed with [expDenote]. This tactic should build a [list (exp * Q)] representing the equations. Remember to give an explicit type annotation when returning a nil list, as in [constr:(@nil (exp * Q))].#</li>#
669 %\item%#<li># Now this final tactic should do the job: 655 %\item%#<li># Now this final tactic should do the job:
656
670 [[ 657 [[
671 Ltac reflectContext := 658 Ltac reflectContext :=
672 let ls := findVarsHyps in 659 let ls := findVarsHyps in
673 repeat match goal with 660 repeat match goal with
674 | [ H : ?e == ?num # ?den |- _ ] => 661 | [ H : ?e == ?num # ?den |- _ ] =>