annotate src/Reflection.v @ 360:e0d91bcf70ec

Pass through old content of Reflection
author Adam Chlipala <adam@chlipala.net>
date Wed, 02 Nov 2011 14:48:25 -0400
parents 3322367e955d
children df17b699a04f
rev   line source
adam@297 1 (* Copyright (c) 2008-2011, Adam Chlipala
adamc@142 2 *
adamc@142 3 * This work is licensed under a
adamc@142 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@142 5 * Unported License.
adamc@142 6 * The license text is available at:
adamc@142 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@142 8 *)
adamc@142 9
adamc@142 10 (* begin hide *)
adamc@142 11 Require Import List.
adamc@142 12
adam@314 13 Require Import CpdtTactics MoreSpecif.
adamc@142 14
adamc@142 15 Set Implicit Arguments.
adamc@142 16 (* end hide *)
adamc@142 17
adamc@142 18
adamc@142 19 (** %\chapter{Proof by Reflection}% *)
adamc@142 20
adam@360 21 (** The last chapter highlighted a very heuristic approach to proving. In this chapter, we will study an alternative technique, %\index{proof by reflection}\textit{%#<i>#proof by reflection#</i>#%}~\cite{reflection}%. We will write, in Gallina, decision procedures with proofs of correctness, and we will appeal to these procedures in writing very short proofs. Such a proof is checked by running the decision procedure. The term %\textit{%#<i>#reflection#</i>#%}% applies because we will need to translate Gallina propositions into values of inductive types representing syntax, so that Gallina programs may analyze them. *)
adamc@142 22
adamc@142 23
adamc@142 24 (** * Proving Evenness *)
adamc@142 25
adamc@142 26 (** Proving that particular natural number constants are even is certainly something we would rather have happen automatically. The Ltac-programming techniques that we learned in the last chapter make it easy to implement such a procedure. *)
adamc@142 27
adamc@142 28 Inductive isEven : nat -> Prop :=
adamc@144 29 | Even_O : isEven O
adamc@144 30 | Even_SS : forall n, isEven n -> isEven (S (S n)).
adamc@142 31
adamc@148 32 (* begin thide *)
adamc@142 33 Ltac prove_even := repeat constructor.
adamc@148 34 (* end thide *)
adamc@142 35
adamc@142 36 Theorem even_256 : isEven 256.
adamc@142 37 prove_even.
adamc@142 38 Qed.
adamc@142 39
adamc@142 40 Print even_256.
adamc@221 41 (** %\vspace{-.15in}% [[
adamc@142 42 even_256 =
adamc@142 43 Even_SS
adamc@142 44 (Even_SS
adamc@142 45 (Even_SS
adamc@142 46 (Even_SS
adamc@221 47
adamc@142 48 ]]
adamc@142 49
adam@360 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 super-linear in the input value. Coq's implicit arguments mechanism is hiding the values given for parameter [n] of [Even_SS], which is why the proof term only appears linear here. Also, proof terms are represented internally as syntax trees, with opportunity for sharing of node representations, but in this chapter we will measure proof term size as simple textual length or as the number of nodes in the term's syntax tree, two measures that are approximately equivalent. Sometimes apparently large proof terms have enough internal sharing that they take up less memory than we expect, but one avoids having to reason about such sharing by ensuring that the size of a sharing-free version of a term is low enough.
adam@360 51
adam@360 52 Superlinear evenness proof terms seem 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.
adamc@142 53
adamc@142 54 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.
adamc@142 55
adamc@142 56 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.
adamc@142 57
adamc@221 58 For this example, we begin by using a type from the [MoreSpecif] module (included in the book source) to write a certified evenness checker. *)
adamc@142 59
adamc@142 60 Print partial.
adamc@221 61 (** %\vspace{-.15in}% [[
adamc@221 62 Inductive partial (P : Prop) : Set := Proved : P -> [P] | Uncertain : [P]
adamc@221 63
adamc@221 64 ]]
adamc@142 65
adamc@221 66 A [partial P] value is an optional proof of [P]. The notation [[P]] stands for [partial P]. *)
adamc@142 67
adamc@221 68 Local Open Scope partial_scope.
adamc@142 69
adamc@142 70 (** 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. *)
adamc@142 71
adamc@148 72 (* begin thide *)
adam@297 73 Definition check_even : forall n : nat, [isEven n].
adamc@142 74 Hint Constructors isEven.
adamc@142 75
adamc@142 76 refine (fix F (n : nat) : [isEven n] :=
adamc@221 77 match n with
adamc@142 78 | 0 => Yes
adamc@142 79 | 1 => No
adamc@142 80 | S (S n') => Reduce (F n')
adamc@142 81 end); auto.
adamc@142 82 Defined.
adamc@142 83
adam@360 84 (** The function [check_even] may be viewed as a %\emph{%#<i>#verified decision procedure#</i>#%}%, because its type guarantees that it never returns [Yes] for inputs that are not even.
adam@360 85
adam@360 86 Now we can use dependent pattern-matching to write a function that performs a surprising feat. When given a [partial P], this function [partialOut] returns a proof of [P] if the [partial] value contains a proof, and it returns a (useless) proof of [True] otherwise. From the standpoint of ML and Haskell programming, it seems impossible to write such a type, but it is trivial with a [return] annotation. *)
adamc@142 87
adamc@142 88 Definition partialOut (P : Prop) (x : [P]) :=
adamc@142 89 match x return (match x with
adamc@142 90 | Proved _ => P
adamc@142 91 | Uncertain => True
adamc@142 92 end) with
adamc@142 93 | Proved pf => pf
adamc@142 94 | Uncertain => I
adamc@142 95 end.
adamc@142 96
adam@289 97 (** It may seem strange to define a function like this. However, it turns out to be very useful in writing a reflective version of our earlier [prove_even] tactic: *)
adamc@142 98
adamc@142 99 Ltac prove_even_reflective :=
adamc@142 100 match goal with
adamc@142 101 | [ |- isEven ?N] => exact (partialOut (check_even N))
adamc@142 102 end.
adamc@148 103 (* end thide *)
adamc@142 104
adam@360 105 (** We identify which natural number we are considering, and we %``%#"#prove#"#%''% its evenness by pulling the proof out of the appropriate [check_even] call. Recall that the %\index{tactics!exact}%[exact] tactic proves a proposition [P] when given a proof term of precisely type [P]. *)
adamc@142 106
adamc@142 107 Theorem even_256' : isEven 256.
adamc@142 108 prove_even_reflective.
adamc@142 109 Qed.
adamc@142 110
adamc@142 111 Print even_256'.
adamc@221 112 (** %\vspace{-.15in}% [[
adamc@142 113 even_256' = partialOut (check_even 256)
adamc@142 114 : isEven 256
adamc@142 115 ]]
adamc@142 116
adam@289 117 We can see a constant wrapper around the object of the proof. For any even number, this form of proof will suffice. The size of the proof term is now linear in the number being checked, containing two repetitions of the unary form of that number, one of which is hidden above within the implicit argument to [partialOut].
adam@289 118
adam@289 119 What happens if we try the tactic with an odd number? *)
adamc@142 120
adamc@142 121 Theorem even_255 : isEven 255.
adam@360 122 (** %\vspace{-.275in}%[[
adamc@142 123 prove_even_reflective.
adam@360 124 ]]
adamc@142 125
adam@360 126 <<
adamc@142 127 User error: No matching clauses for match goal
adam@360 128 >>
adamc@142 129
adamc@142 130 Thankfully, the tactic fails. To see more precisely what goes wrong, we can run manually the body of the [match].
adamc@142 131
adam@360 132 %\vspace{-.15in}%[[
adamc@142 133 exact (partialOut (check_even 255)).
adam@360 134 ]]
adamc@142 135
adam@360 136 <<
adamc@142 137 Error: The term "partialOut (check_even 255)" has type
adamc@142 138 "match check_even 255 with
adamc@142 139 | Yes => isEven 255
adamc@142 140 | No => True
adamc@142 141 end" while it is expected to have type "isEven 255"
adam@360 142 >>
adamc@142 143
adam@360 144 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]. *)
adamc@221 145
adamc@142 146 Abort.
adamc@143 147
adam@360 148 (** Our tactic [prove_even_reflective] is reflective because it performs a proof search process (a trivial one, in this case) wholly within Gallina, where the only use of Ltac is to translate a goal into an appropriate use of [check_even]. *)
adamc@143 149
adam@360 150
adam@360 151 (** * Reifying the Syntax of a Trivial Tautology Language *)
adamc@143 152
adamc@143 153 (** We might also like to have reflective proofs of trivial tautologies like this one: *)
adamc@143 154
adamc@143 155 Theorem true_galore : (True /\ True) -> (True \/ (True /\ (True -> True))).
adamc@143 156 tauto.
adamc@143 157 Qed.
adamc@143 158
adamc@143 159 Print true_galore.
adamc@221 160 (** %\vspace{-.15in}% [[
adamc@143 161 true_galore =
adamc@143 162 fun H : True /\ True =>
adamc@143 163 and_ind (fun _ _ : True => or_introl (True /\ (True -> True)) I) H
adamc@143 164 : True /\ True -> True \/ True /\ (True -> True)
adamc@221 165
adamc@143 166 ]]
adamc@143 167
adamc@143 168 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.
adamc@143 169
adam@360 170 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 %\index{reification}\textit{%#<i>#reify#</i>#%}% [Prop] into some type that we %\textit{%#<i>#can#</i>#%}% analyze. This inductive type is a good candidate: *)
adamc@143 171
adamc@148 172 (* begin thide *)
adamc@143 173 Inductive taut : Set :=
adamc@143 174 | TautTrue : taut
adamc@143 175 | TautAnd : taut -> taut -> taut
adamc@143 176 | TautOr : taut -> taut -> taut
adamc@143 177 | TautImp : taut -> taut -> taut.
adamc@143 178
adam@360 179 (** We write a recursive function to %\emph{%#<i>#reflect#</i>#%}% this syntax back to [Prop]. Such functions are also called %\index{interpretation function}\emph{%#<i>#interpretation functions#</i>#%}%, and have used them in previous examples to give semantics to small programming languages. *)
adamc@143 180
adamc@143 181 Fixpoint tautDenote (t : taut) : Prop :=
adamc@143 182 match t with
adamc@143 183 | TautTrue => True
adamc@143 184 | TautAnd t1 t2 => tautDenote t1 /\ tautDenote t2
adamc@143 185 | TautOr t1 t2 => tautDenote t1 \/ tautDenote t2
adamc@143 186 | TautImp t1 t2 => tautDenote t1 -> tautDenote t2
adamc@143 187 end.
adamc@143 188
adamc@143 189 (** It is easy to prove that every formula in the range of [tautDenote] is true. *)
adamc@143 190
adamc@143 191 Theorem tautTrue : forall t, tautDenote t.
adamc@143 192 induction t; crush.
adamc@143 193 Qed.
adamc@143 194
adam@360 195 (** To use [tautTrue] to prove particular formulas, we need to implement the syntax reification process. A recursive Ltac function does the job. *)
adamc@143 196
adam@360 197 Ltac tautReify P :=
adamc@143 198 match P with
adamc@143 199 | True => TautTrue
adamc@143 200 | ?P1 /\ ?P2 =>
adam@360 201 let t1 := tautReify P1 in
adam@360 202 let t2 := tautReify P2 in
adamc@143 203 constr:(TautAnd t1 t2)
adamc@143 204 | ?P1 \/ ?P2 =>
adam@360 205 let t1 := tautReify P1 in
adam@360 206 let t2 := tautReify P2 in
adamc@143 207 constr:(TautOr t1 t2)
adamc@143 208 | ?P1 -> ?P2 =>
adam@360 209 let t1 := tautReify P1 in
adam@360 210 let t2 := tautReify P2 in
adamc@143 211 constr:(TautImp t1 t2)
adamc@143 212 end.
adamc@143 213
adam@360 214 (** With [tautReify] available, it is easy to finish our reflective tactic. We look at the goal formula, reflect it, and apply [tautTrue] to the reflected formula. *)
adamc@143 215
adamc@143 216 Ltac obvious :=
adamc@143 217 match goal with
adamc@143 218 | [ |- ?P ] =>
adam@360 219 let t := tautReify P in
adamc@143 220 exact (tautTrue t)
adamc@143 221 end.
adamc@143 222
adamc@143 223 (** We can verify that [obvious] solves our original example, with a proof term that does not mention details of the proof. *)
adamc@148 224 (* end thide *)
adamc@143 225
adamc@143 226 Theorem true_galore' : (True /\ True) -> (True \/ (True /\ (True -> True))).
adamc@143 227 obvious.
adamc@143 228 Qed.
adamc@143 229
adamc@143 230 Print true_galore'.
adamc@143 231
adamc@221 232 (** %\vspace{-.15in}% [[
adamc@143 233 true_galore' =
adamc@143 234 tautTrue
adamc@143 235 (TautImp (TautAnd TautTrue TautTrue)
adamc@143 236 (TautOr TautTrue (TautAnd TautTrue (TautImp TautTrue TautTrue))))
adamc@143 237 : True /\ True -> True \/ True /\ (True -> True)
adamc@143 238 ]]
adamc@143 239
adam@360 240 It is worth considering how the reflective tactic improves on a pure-Ltac implementation. The formula reification 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.
adam@360 241
adam@360 242 It may also be worth pointing out that our previous example of evenness testing used a function [partialOut] for sound handling of input goals that the verified decision procedure fails to prove. Here, we prove that our procedure [tautTrue] (recall that an inductive proof may be viewed as a recursive procedure) is able to prove any goal representable in [taut], so no extra step is necessary. *)
adamc@144 243
adamc@144 244
adamc@145 245 (** * A Monoid Expression Simplifier *)
adamc@145 246
adam@289 247 (** Proof by reflection does not require encoding of all of the syntax in a goal. We can insert %``%#"#variables#"#%''% in our syntax types to allow injection of arbitrary pieces, even if we cannot apply specialized reasoning to them. In this section, we explore that possibility by writing a tactic for normalizing monoid equations. *)
adamc@146 248
adamc@145 249 Section monoid.
adamc@145 250 Variable A : Set.
adamc@145 251 Variable e : A.
adamc@145 252 Variable f : A -> A -> A.
adamc@145 253
adamc@145 254 Infix "+" := f.
adamc@145 255
adamc@145 256 Hypothesis assoc : forall a b c, (a + b) + c = a + (b + c).
adamc@145 257 Hypothesis identl : forall a, e + a = a.
adamc@145 258 Hypothesis identr : forall a, a + e = a.
adamc@145 259
adamc@146 260 (** We add variables and hypotheses characterizing an arbitrary instance of the algebraic structure of monoids. We have an associative binary operator and an identity element for it.
adamc@146 261
adam@289 262 It is easy to define an expression tree type for monoid expressions. A [Var] constructor is a %``%#"#catch-all#"#%''% case for subexpressions that we cannot model. These subexpressions could be actual Gallina variables, or they could just use functions that our tactic is unable to understand. *)
adamc@146 263
adamc@148 264 (* begin thide *)
adamc@145 265 Inductive mexp : Set :=
adamc@145 266 | Ident : mexp
adamc@145 267 | Var : A -> mexp
adamc@145 268 | Op : mexp -> mexp -> mexp.
adamc@145 269
adam@360 270 (** Next, we write an interpretation function. *)
adamc@146 271
adamc@145 272 Fixpoint mdenote (me : mexp) : A :=
adamc@145 273 match me with
adamc@145 274 | Ident => e
adamc@145 275 | Var v => v
adamc@145 276 | Op me1 me2 => mdenote me1 + mdenote me2
adamc@145 277 end.
adamc@145 278
adamc@146 279 (** We will normalize expressions by flattening them into lists, via associativity, so it is helpful to have a denotation function for lists of monoid values. *)
adamc@146 280
adamc@145 281 Fixpoint mldenote (ls : list A) : A :=
adamc@145 282 match ls with
adamc@145 283 | nil => e
adamc@145 284 | x :: ls' => x + mldenote ls'
adamc@145 285 end.
adamc@145 286
adamc@146 287 (** The flattening function itself is easy to implement. *)
adamc@146 288
adamc@145 289 Fixpoint flatten (me : mexp) : list A :=
adamc@145 290 match me with
adamc@145 291 | Ident => nil
adamc@145 292 | Var x => x :: nil
adamc@145 293 | Op me1 me2 => flatten me1 ++ flatten me2
adamc@145 294 end.
adamc@145 295
adamc@146 296 (** [flatten] has a straightforward correctness proof in terms of our [denote] functions. *)
adamc@146 297
adamc@146 298 Lemma flatten_correct' : forall ml2 ml1,
adamc@146 299 mldenote ml1 + mldenote ml2 = mldenote (ml1 ++ ml2).
adamc@145 300 induction ml1; crush.
adamc@145 301 Qed.
adamc@145 302
adamc@145 303 Theorem flatten_correct : forall me, mdenote me = mldenote (flatten me).
adamc@145 304 Hint Resolve flatten_correct'.
adamc@145 305
adamc@145 306 induction me; crush.
adamc@145 307 Qed.
adamc@145 308
adamc@146 309 (** Now it is easy to prove a theorem that will be the main tool behind our simplification tactic. *)
adamc@146 310
adamc@146 311 Theorem monoid_reflect : forall me1 me2,
adamc@146 312 mldenote (flatten me1) = mldenote (flatten me2)
adamc@146 313 -> mdenote me1 = mdenote me2.
adamc@145 314 intros; repeat rewrite flatten_correct; assumption.
adamc@145 315 Qed.
adamc@145 316
adam@360 317 (** We implement reification into the [mexp] type. *)
adamc@146 318
adam@360 319 Ltac reify me :=
adamc@146 320 match me with
adamc@145 321 | e => Ident
adamc@146 322 | ?me1 + ?me2 =>
adam@360 323 let r1 := reify me1 in
adam@360 324 let r2 := reify me2 in
adamc@145 325 constr:(Op r1 r2)
adamc@146 326 | _ => constr:(Var me)
adamc@145 327 end.
adamc@145 328
adam@360 329 (** The final [monoid] tactic works on goals that equate two monoid terms. We reify each and change the goal to refer to the reified versions, finishing off by applying [monoid_reflect] and simplifying uses of [mldenote]. Recall that the %\index{tactics!change}%[change] tactic replaces a conclusion formula with another that is definitionally equal to it. *)
adamc@146 330
adamc@145 331 Ltac monoid :=
adamc@145 332 match goal with
adamc@146 333 | [ |- ?me1 = ?me2 ] =>
adam@360 334 let r1 := reify me1 in
adam@360 335 let r2 := reify me2 in
adamc@145 336 change (mdenote r1 = mdenote r2);
adam@360 337 apply monoid_reflect; simpl
adamc@145 338 end.
adamc@145 339
adamc@146 340 (** We can make short work of theorems like this one: *)
adamc@146 341
adamc@148 342 (* end thide *)
adamc@148 343
adamc@145 344 Theorem t1 : forall a b c d, a + b + c + d = a + (b + c) + d.
adamc@146 345 intros; monoid.
adamc@146 346 (** [[
adamc@146 347 ============================
adamc@146 348 a + (b + (c + (d + e))) = a + (b + (c + (d + e)))
adamc@221 349
adamc@146 350 ]]
adamc@146 351
adam@360 352 Our tactic has canonicalized both sides of the equality, such that we can finish the proof by reflexivity. *)
adamc@146 353
adamc@145 354 reflexivity.
adamc@145 355 Qed.
adamc@146 356
adamc@146 357 (** It is interesting to look at the form of the proof. *)
adamc@146 358
adamc@146 359 Print t1.
adamc@221 360 (** %\vspace{-.15in}% [[
adamc@146 361 t1 =
adamc@146 362 fun a b c d : A =>
adamc@146 363 monoid_reflect (Op (Op (Op (Var a) (Var b)) (Var c)) (Var d))
adamc@146 364 (Op (Op (Var a) (Op (Var b) (Var c))) (Var d))
adamc@146 365 (refl_equal (a + (b + (c + (d + e)))))
adamc@146 366 : forall a b c d : A, a + b + c + d = a + (b + c) + d
adamc@146 367 ]]
adamc@146 368
adam@360 369 The proof term contains only restatements of the equality operands in reified form, followed by a use of reflexivity on the shared canonical form. *)
adamc@221 370
adamc@145 371 End monoid.
adamc@145 372
adam@360 373 (** Extensions of this basic approach are used in the implementations of the %\index{tactics!ring}%[ring] and %\index{tactics!field}%[field] tactics that come packaged with Coq. *)
adamc@146 374
adamc@145 375
adamc@144 376 (** * A Smarter Tautology Solver *)
adamc@144 377
adam@360 378 (** 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.
adamc@147 379
adam@360 380 To arrive at a nice implementation satisfying these criteria, we introduce the %\index{tactics!quote}%[quote] tactic and its associated library. *)
adamc@147 381
adamc@144 382 Require Import Quote.
adamc@144 383
adamc@148 384 (* begin thide *)
adamc@144 385 Inductive formula : Set :=
adamc@144 386 | Atomic : index -> formula
adamc@144 387 | Truth : formula
adamc@144 388 | Falsehood : formula
adamc@144 389 | And : formula -> formula -> formula
adamc@144 390 | Or : formula -> formula -> formula
adamc@144 391 | Imp : formula -> formula -> formula.
adamc@144 392
adam@360 393 (** The type %\index{Gallina terms!index}%[index] comes from the [Quote] library and represents a countable variable type. The rest of [formula]'s definition should be old hat by now.
adamc@147 394
adamc@147 395 The [quote] tactic will implement injection from [Prop] into [formula] for us, but it is not quite as smart as we might like. In particular, it interprets implications incorrectly, so we will need to declare a wrapper definition for implication, as we did in the last chapter. *)
adamc@144 396
adamc@144 397 Definition imp (P1 P2 : Prop) := P1 -> P2.
adamc@144 398 Infix "-->" := imp (no associativity, at level 95).
adamc@144 399
adamc@147 400 (** Now we can define our denotation function. *)
adamc@147 401
adamc@147 402 Definition asgn := varmap Prop.
adamc@147 403
adamc@144 404 Fixpoint formulaDenote (atomics : asgn) (f : formula) : Prop :=
adamc@144 405 match f with
adamc@144 406 | Atomic v => varmap_find False v atomics
adamc@144 407 | Truth => True
adamc@144 408 | Falsehood => False
adamc@144 409 | And f1 f2 => formulaDenote atomics f1 /\ formulaDenote atomics f2
adamc@144 410 | Or f1 f2 => formulaDenote atomics f1 \/ formulaDenote atomics f2
adamc@144 411 | Imp f1 f2 => formulaDenote atomics f1 --> formulaDenote atomics f2
adamc@144 412 end.
adamc@144 413
adam@360 414 (** The %\index{Gallina terms!varmap}%[varmap] type family implements maps from [index] values. In this case, we define an assignment as a map from variables to [Prop]s. Our reifier [formulaDenote] works with an assignment, and we use the [varmap_find] function to consult the assignment in the [Atomic] case. The first argument to [varmap_find] is a default value, in case the variable is not found. *)
adamc@147 415
adamc@144 416 Section my_tauto.
adamc@144 417 Variable atomics : asgn.
adamc@144 418
adamc@144 419 Definition holds (v : index) := varmap_find False v atomics.
adamc@144 420
adamc@147 421 (** We define some shorthand for a particular variable being true, and now we are ready to define some helpful functions based on the [ListSet] module of the standard library, which (unsurprisingly) presents a view of lists as sets. *)
adamc@147 422
adamc@144 423 Require Import ListSet.
adamc@144 424
adamc@144 425 Definition index_eq : forall x y : index, {x = y} + {x <> y}.
adamc@144 426 decide equality.
adamc@144 427 Defined.
adamc@144 428
adamc@144 429 Definition add (s : set index) (v : index) := set_add index_eq v s.
adamc@147 430
adamc@221 431 Definition In_dec : forall v (s : set index), {In v s} + {~ In v s}.
adamc@221 432 Local Open Scope specif_scope.
adamc@144 433
adamc@221 434 intro; refine (fix F (s : set index) : {In v s} + {~ In v s} :=
adamc@221 435 match s with
adamc@144 436 | nil => No
adamc@144 437 | v' :: s' => index_eq v' v || F s'
adamc@144 438 end); crush.
adamc@144 439 Defined.
adamc@144 440
adamc@147 441 (** We define what it means for all members of an index set to represent true propositions, and we prove some lemmas about this notion. *)
adamc@147 442
adamc@144 443 Fixpoint allTrue (s : set index) : Prop :=
adamc@144 444 match s with
adamc@144 445 | nil => True
adamc@144 446 | v :: s' => holds v /\ allTrue s'
adamc@144 447 end.
adamc@144 448
adamc@144 449 Theorem allTrue_add : forall v s,
adamc@144 450 allTrue s
adamc@144 451 -> holds v
adamc@144 452 -> allTrue (add s v).
adamc@144 453 induction s; crush;
adamc@144 454 match goal with
adamc@144 455 | [ |- context[if ?E then _ else _] ] => destruct E
adamc@144 456 end; crush.
adamc@144 457 Qed.
adamc@144 458
adamc@144 459 Theorem allTrue_In : forall v s,
adamc@144 460 allTrue s
adamc@144 461 -> set_In v s
adamc@144 462 -> varmap_find False v atomics.
adamc@144 463 induction s; crush.
adamc@144 464 Qed.
adamc@144 465
adamc@144 466 Hint Resolve allTrue_add allTrue_In.
adamc@144 467
adamc@221 468 Local Open Scope partial_scope.
adamc@144 469
adam@353 470 (** 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]. *)
adamc@147 471
adam@297 472 Definition forward : forall (f : formula) (known : set index) (hyp : formula)
adam@297 473 (cont : forall known', [allTrue known' -> formulaDenote atomics f]),
adam@297 474 [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f].
adamc@144 475 refine (fix F (f : formula) (known : set index) (hyp : formula)
adamc@221 476 (cont : forall known', [allTrue known' -> formulaDenote atomics f])
adamc@144 477 : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f] :=
adamc@221 478 match hyp with
adamc@144 479 | Atomic v => Reduce (cont (add known v))
adamc@144 480 | Truth => Reduce (cont known)
adamc@144 481 | Falsehood => Yes
adamc@144 482 | And h1 h2 =>
adamc@144 483 Reduce (F (Imp h2 f) known h1 (fun known' =>
adamc@144 484 Reduce (F f known' h2 cont)))
adamc@144 485 | Or h1 h2 => F f known h1 cont && F f known h2 cont
adamc@144 486 | Imp _ _ => Reduce (cont known)
adamc@144 487 end); crush.
adamc@144 488 Defined.
adamc@144 489
adamc@147 490 (** A [backward] function implements analysis of the final goal. It calls [forward] to handle implications. *)
adamc@147 491
adam@297 492 Definition backward : forall (known : set index) (f : formula),
adam@297 493 [allTrue known -> formulaDenote atomics f].
adamc@221 494 refine (fix F (known : set index) (f : formula)
adamc@221 495 : [allTrue known -> formulaDenote atomics f] :=
adamc@221 496 match f with
adamc@144 497 | Atomic v => Reduce (In_dec v known)
adamc@144 498 | Truth => Yes
adamc@144 499 | Falsehood => No
adamc@144 500 | And f1 f2 => F known f1 && F known f2
adamc@144 501 | Or f1 f2 => F known f1 || F known f2
adamc@144 502 | Imp f1 f2 => forward f2 known f1 (fun known' => F known' f2)
adamc@144 503 end); crush; eauto.
adamc@144 504 Defined.
adamc@144 505
adamc@147 506 (** A simple wrapper around [backward] gives us the usual type of a partial decision procedure. *)
adamc@147 507
adam@297 508 Definition my_tauto : forall f : formula, [formulaDenote atomics f].
adamc@144 509 intro; refine (Reduce (backward nil f)); crush.
adamc@144 510 Defined.
adamc@144 511 End my_tauto.
adamc@144 512
adam@360 513 (** Our final tactic implementation is now fairly straightforward. First, we [intro] all quantifiers that do not bind [Prop]s. Then we call the [quote] tactic, which implements the reification for us. Finally, we are able to construct an exact proof via [partialOut] and the [my_tauto] Gallina function. *)
adamc@147 514
adamc@144 515 Ltac my_tauto :=
adamc@144 516 repeat match goal with
adamc@144 517 | [ |- forall x : ?P, _ ] =>
adamc@144 518 match type of P with
adamc@144 519 | Prop => fail 1
adamc@144 520 | _ => intro
adamc@144 521 end
adamc@144 522 end;
adamc@144 523 quote formulaDenote;
adamc@144 524 match goal with
adamc@144 525 | [ |- formulaDenote ?m ?f ] => exact (partialOut (my_tauto m f))
adamc@144 526 end.
adamc@148 527 (* end thide *)
adamc@144 528
adamc@147 529 (** A few examples demonstrate how the tactic works. *)
adamc@147 530
adamc@144 531 Theorem mt1 : True.
adamc@144 532 my_tauto.
adamc@144 533 Qed.
adamc@144 534
adamc@144 535 Print mt1.
adamc@221 536 (** %\vspace{-.15in}% [[
adamc@147 537 mt1 = partialOut (my_tauto (Empty_vm Prop) Truth)
adamc@147 538 : True
adamc@147 539 ]]
adamc@147 540
adamc@147 541 We see [my_tauto] applied with an empty [varmap], since every subformula is handled by [formulaDenote]. *)
adamc@144 542
adamc@144 543 Theorem mt2 : forall x y : nat, x = y --> x = y.
adamc@144 544 my_tauto.
adamc@144 545 Qed.
adamc@144 546
adamc@144 547 Print mt2.
adamc@221 548 (** %\vspace{-.15in}% [[
adamc@147 549 mt2 =
adamc@147 550 fun x y : nat =>
adamc@147 551 partialOut
adamc@147 552 (my_tauto (Node_vm (x = y) (Empty_vm Prop) (Empty_vm Prop))
adamc@147 553 (Imp (Atomic End_idx) (Atomic End_idx)))
adamc@147 554 : forall x y : nat, x = y --> x = y
adamc@147 555 ]]
adamc@147 556
adamc@147 557 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. *)
adamc@144 558
adamc@144 559 Theorem mt3 : forall x y z,
adamc@144 560 (x < y /\ y > z) \/ (y > z /\ x < S y)
adamc@144 561 --> y > z /\ (x < y \/ x < S y).
adamc@144 562 my_tauto.
adamc@144 563 Qed.
adamc@144 564
adamc@144 565 Print mt3.
adamc@221 566 (** %\vspace{-.15in}% [[
adamc@147 567 fun x y z : nat =>
adamc@147 568 partialOut
adamc@147 569 (my_tauto
adamc@147 570 (Node_vm (x < S y) (Node_vm (x < y) (Empty_vm Prop) (Empty_vm Prop))
adamc@147 571 (Node_vm (y > z) (Empty_vm Prop) (Empty_vm Prop)))
adamc@147 572 (Imp
adamc@147 573 (Or (And (Atomic (Left_idx End_idx)) (Atomic (Right_idx End_idx)))
adamc@147 574 (And (Atomic (Right_idx End_idx)) (Atomic End_idx)))
adamc@147 575 (And (Atomic (Right_idx End_idx))
adamc@147 576 (Or (Atomic (Left_idx End_idx)) (Atomic End_idx)))))
adamc@147 577 : forall x y z : nat,
adamc@147 578 x < y /\ y > z \/ y > z /\ x < S y --> y > z /\ (x < y \/ x < S y)
adamc@147 579 ]]
adamc@147 580
adamc@147 581 Our goal contained three distinct atomic formulas, and we see that a three-element [varmap] is generated.
adamc@147 582
adamc@147 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. *)
adamc@144 584
adamc@144 585 Theorem mt4 : True /\ True /\ True /\ True /\ True /\ True /\ False --> False.
adamc@144 586 my_tauto.
adamc@144 587 Qed.
adamc@144 588
adamc@144 589 Print mt4.
adamc@221 590 (** %\vspace{-.15in}% [[
adamc@147 591 mt4 =
adamc@147 592 partialOut
adamc@147 593 (my_tauto (Empty_vm Prop)
adamc@147 594 (Imp
adamc@147 595 (And Truth
adamc@147 596 (And Truth
adamc@147 597 (And Truth (And Truth (And Truth (And Truth Falsehood))))))
adamc@147 598 Falsehood))
adamc@147 599 : True /\ True /\ True /\ True /\ True /\ True /\ False --> False
adam@302 600 ]]
adam@302 601 *)
adamc@144 602
adamc@144 603 Theorem mt4' : True /\ True /\ True /\ True /\ True /\ True /\ False -> False.
adamc@144 604 tauto.
adamc@144 605 Qed.
adamc@144 606
adamc@144 607 Print mt4'.
adamc@221 608 (** %\vspace{-.15in}% [[
adamc@147 609 mt4' =
adamc@147 610 fun H : True /\ True /\ True /\ True /\ True /\ True /\ False =>
adamc@147 611 and_ind
adamc@147 612 (fun (_ : True) (H1 : True /\ True /\ True /\ True /\ True /\ False) =>
adamc@147 613 and_ind
adamc@147 614 (fun (_ : True) (H3 : True /\ True /\ True /\ True /\ False) =>
adamc@147 615 and_ind
adamc@147 616 (fun (_ : True) (H5 : True /\ True /\ True /\ False) =>
adamc@147 617 and_ind
adamc@147 618 (fun (_ : True) (H7 : True /\ True /\ False) =>
adamc@147 619 and_ind
adamc@147 620 (fun (_ : True) (H9 : True /\ False) =>
adamc@147 621 and_ind (fun (_ : True) (H11 : False) => False_ind False H11)
adamc@147 622 H9) H7) H5) H3) H1) H
adamc@147 623 : True /\ True /\ True /\ True /\ True /\ True /\ False -> False
adam@302 624 ]]
adam@360 625
adam@360 626 The traditional [tauto] tactic introduces a quadratic blow-up in the size of the proof term, whereas proofs produced by [my_tauto] always have linear size. *)
adamc@147 627
adamc@149 628
adamc@149 629 (** * Exercises *)
adamc@149 630
adamc@221 631 (** remove printing * *)
adamc@221 632
adamc@149 633 (** %\begin{enumerate}%#<ol>#
adamc@149 634
adamc@149 635 %\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.
adamc@149 636
adamc@221 637 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:
adamc@149 638 [[
adamc@149 639 Theorem t2 : forall x y z, (2 # 1) * (x - (3 # 2) * y) == 15 # 1
adamc@149 640 -> z + (8 # 1) * x == 20 # 1
adamc@149 641 -> (-6 # 2) * y + (10 # 1) * x + z == 35 # 1.
adam@360 642 intros; reifyContext; assumption.
adamc@149 643 Qed.
adamc@205 644 ]]
adamc@205 645
adam@360 646 Your solution can work in any way that involves reifying 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 [==].
adamc@149 647
adamc@149 648 %\begin{enumerate}%#<ol>#
adamc@221 649 %\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>#
adamc@149 650 %\item%#<li># Define a function [lookup] for reading an element out of a list of rationals, by its position in the list.#</li>#
adamc@149 651 %\item%#<li># Define a function [expDenote] that translates [exp]s, along with lists of rationals representing variable values, to [Q].#</li>#
adamc@149 652 %\item%#<li># Define a recursive function [eqsDenote] over [list (exp * Q)], characterizing when all of the equations are true.#</li>#
adamc@149 653 %\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>#
adamc@151 654 %\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>#
adamc@149 655 %\item%#<li># Write a recursive function [linearizeEqs : list (exp * Q) -> option (lhs * Q)]. This function linearizes all of the equations in the list in turn, building up the sum of the equations. It returns [None] if the linearization of any constituent equation fails.#</li>#
adamc@149 656 %\item%#<li># Define a denotation function for [lhs].#</li>#
adamc@149 657 %\item%#<li># Prove that, when [exp] linearization succeeds on constant [k] and expression [e], the linearized version has the same meaning as [k * e].#</li>#
adamc@149 658 %\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>#
adamc@149 659 %\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>#
adam@360 660 %\item%#<li># Write a tactic [reify] to reify a [Q] expression into [exp], with respect to a given list of variable values.#</li>#
adam@360 661 %\item%#<li># Write a tactic [reifyEqs] to reify 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>#
adamc@149 662 %\item%#<li># Now this final tactic should do the job:
adamc@149 663 [[
adam@360 664 Ltac reifyContext :=
adamc@149 665 let ls := findVarsHyps in
adamc@149 666 repeat match goal with
adamc@149 667 | [ H : ?e == ?num # ?den |- _ ] =>
adam@360 668 let r := reify ls e in
adamc@149 669 change (expDenote ls r == num # den) in H;
adamc@149 670 generalize H
adamc@149 671 end;
adamc@149 672 match goal with
adam@360 673 | [ |- ?g ] => let re := reifyEqs g in
adamc@149 674 intros;
adamc@149 675 let H := fresh "H" in
adamc@149 676 assert (H : eqsDenote ls re); [ simpl in *; tauto
adamc@149 677 | repeat match goal with
adamc@149 678 | [ H : expDenote _ _ == _ |- _ ] => clear H
adamc@149 679 end;
adamc@149 680 generalize (linearizeEqsCorrect ls re H); clear H; simpl;
adamc@149 681 match goal with
adamc@149 682 | [ |- ?X == ?Y -> _ ] =>
adamc@149 683 ring_simplify X Y; intro
adamc@149 684 end ]
adamc@149 685 end.
adamc@205 686 ]]
adamc@205 687
adamc@149 688 #</ol>#%\end{enumerate}%
adamc@149 689 #</li>#
adamc@149 690
adamc@149 691 #</ol>#%\end{enumerate}% *)