Reflection.v 29.1 KB
Newer Older
1
(* Copyright (c) 2008-2010, Adam Chlipala
Adam Chlipala's avatar
Adam Chlipala committed
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
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)

(* begin hide *)
Require Import List.

Require Import Tactics MoreSpecif.

Set Implicit Arguments.
(* end hide *)


(** %\chapter{Proof by Reflection}% *)

(** The last chapter highlighted a very heuristic approach to proving.  In this chapter, we will study an alternative technique, %\textit{%#<i>#proof by reflection#</i>#%}%.  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. *)


(** * Proving Evenness *)

(** 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. *)

Inductive isEven : nat -> Prop :=
Adam Chlipala's avatar
Adam Chlipala committed
29 30
| Even_O : isEven O
| Even_SS : forall n, isEven n -> isEven (S (S n)).
Adam Chlipala's avatar
Adam Chlipala committed
31

Adam Chlipala's avatar
Adam Chlipala committed
32
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
33
Ltac prove_even := repeat constructor.
Adam Chlipala's avatar
Adam Chlipala committed
34
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
35 36 37 38 39 40

Theorem even_256 : isEven 256.
  prove_even.
Qed.

Print even_256.
Adam Chlipala's avatar
Adam Chlipala committed
41
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
42 43 44 45 46
even_256 = 
Even_SS
  (Even_SS
     (Even_SS
        (Even_SS
Adam Chlipala's avatar
Adam Chlipala committed
47
 
Adam Chlipala's avatar
Adam Chlipala committed
48 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 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.)  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.
Adam Chlipala's avatar
Adam Chlipala committed
51 52 53 54 55

    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.

    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.

Adam Chlipala's avatar
Adam Chlipala committed
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. *)
Adam Chlipala's avatar
Adam Chlipala committed
57 58

Print partial.
Adam Chlipala's avatar
Adam Chlipala committed
59
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
60
Inductive partial (P : Prop) : Set :=  Proved : P -> [P] | Uncertain : [P]
Adam Chlipala's avatar
Adam Chlipala committed
61 62
 
    ]]
Adam Chlipala's avatar
Adam Chlipala committed
63

Adam Chlipala's avatar
Adam Chlipala committed
64
    A [partial P] value is an optional proof of [P]. The notation [[P]] stands for [partial P]. *)
Adam Chlipala's avatar
Adam Chlipala committed
65

Adam Chlipala's avatar
Adam Chlipala committed
66
Local Open Scope partial_scope.
Adam Chlipala's avatar
Adam Chlipala committed
67 68 69

(** 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. *)

Adam Chlipala's avatar
Adam Chlipala committed
70
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
71 72 73 74
Definition check_even (n : nat) : [isEven n].
  Hint Constructors isEven.

  refine (fix F (n : nat) : [isEven n] :=
Adam Chlipala's avatar
Adam Chlipala committed
75
    match n with
Adam Chlipala's avatar
Adam Chlipala committed
76 77 78 79 80 81 82 83 84 85 86 87 88 89 90 91 92
      | 0 => Yes
      | 1 => No
      | S (S n') => Reduce (F n')
    end); auto.
Defined.

(** 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. *)

Definition partialOut (P : Prop) (x : [P]) :=
  match x return (match x with
                    | Proved _ => P
                    | Uncertain => True
                  end) with
    | Proved pf => pf
    | Uncertain => I
  end.

93
(** 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: *)
Adam Chlipala's avatar
Adam Chlipala committed
94 95 96 97 98

Ltac prove_even_reflective :=
  match goal with
    | [ |- isEven ?N] => exact (partialOut (check_even N))
  end.
Adam Chlipala's avatar
Adam Chlipala committed
99
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
100

101
(** We identify which natural number we are considering, and we %``%#"#prove#"#%''% its evenness by pulling the proof out of the appropriate [check_even] call. *)
Adam Chlipala's avatar
Adam Chlipala committed
102 103 104 105 106 107

Theorem even_256' : isEven 256.
  prove_even_reflective.
Qed.

Print even_256'.
Adam Chlipala's avatar
Adam Chlipala committed
108
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
109 110
even_256' = partialOut (check_even 256)
     : isEven 256
Adam Chlipala's avatar
Adam Chlipala committed
111
 
Adam Chlipala's avatar
Adam Chlipala committed
112 113
    ]]

114 115 116
    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].

    What happens if we try the tactic with an odd number? *)
Adam Chlipala's avatar
Adam Chlipala committed
117 118 119 120 121 122

Theorem even_255 : isEven 255.
  (** [[
  prove_even_reflective.

User error: No matching clauses for match goal
Adam Chlipala's avatar
Adam Chlipala committed
123
 
Adam Chlipala's avatar
Adam Chlipala committed
124 125 126 127 128 129 130 131 132 133 134 135
  ]]

  Thankfully, the tactic fails.  To see more precisely what goes wrong, we can run manually the body of the [match].

  [[
  exact (partialOut (check_even 255)).

  Error: The term "partialOut (check_even 255)" has type
 "match check_even 255 with
  | Yes => isEven 255
  | No => True
  end" while it is expected to have type "isEven 255"
Adam Chlipala's avatar
Adam Chlipala committed
136
 
Adam Chlipala's avatar
Adam Chlipala committed
137 138 139
  ]]

  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]. *)
Adam Chlipala's avatar
Adam Chlipala committed
140

Adam Chlipala's avatar
Adam Chlipala committed
141
Abort.
Adam Chlipala's avatar
Adam Chlipala committed
142 143 144 145 146 147 148 149 150 151 152


(** * Reflecting the Syntax of a Trivial Tautology Language *)

(** We might also like to have reflective proofs of trivial tautologies like this one: *)

Theorem true_galore : (True /\ True) -> (True \/ (True /\ (True -> True))).
  tauto.
Qed.

Print true_galore.
Adam Chlipala's avatar
Adam Chlipala committed
153
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
154 155 156 157
true_galore = 
fun H : True /\ True =>
and_ind (fun _ _ : True => or_introl (True /\ (True -> True)) I) H
     : True /\ True -> True \/ True /\ (True -> True)
Adam Chlipala's avatar
Adam Chlipala committed
158
 
Adam Chlipala's avatar
Adam Chlipala committed
159 160 161 162
    ]]

    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.

163
   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: *)
Adam Chlipala's avatar
Adam Chlipala committed
164

Adam Chlipala's avatar
Adam Chlipala committed
165
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
166 167 168 169 170 171
Inductive taut : Set :=
| TautTrue : taut
| TautAnd : taut -> taut -> taut
| TautOr : taut -> taut -> taut
| TautImp : taut -> taut -> taut.

172
(** We write a recursive function to %``%#"#unreflect#"#%''% this syntax back to [Prop]. *)
Adam Chlipala's avatar
Adam Chlipala committed
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

Fixpoint tautDenote (t : taut) : Prop :=
  match t with
    | TautTrue => True
    | TautAnd t1 t2 => tautDenote t1 /\ tautDenote t2
    | TautOr t1 t2 => tautDenote t1 \/ tautDenote t2
    | TautImp t1 t2 => tautDenote t1 -> tautDenote t2
  end.

(** It is easy to prove that every formula in the range of [tautDenote] is true. *)

Theorem tautTrue : forall t, tautDenote t.
  induction t; crush.
Qed.

(** To use [tautTrue] to prove particular formulas, we need to implement the syntax reflection process.  A recursive Ltac function does the job. *)

Ltac tautReflect P :=
  match P with
    | True => TautTrue
    | ?P1 /\ ?P2 =>
      let t1 := tautReflect P1 in
      let t2 := tautReflect P2 in
        constr:(TautAnd t1 t2)
    | ?P1 \/ ?P2 =>
      let t1 := tautReflect P1 in
      let t2 := tautReflect P2 in
        constr:(TautOr t1 t2)
    | ?P1 -> ?P2 =>
      let t1 := tautReflect P1 in
      let t2 := tautReflect P2 in
        constr:(TautImp t1 t2)
  end.

(** With [tautReflect] available, it is easy to finish our reflective tactic.  We look at the goal formula, reflect it, and apply [tautTrue] to the reflected formula. *)

Ltac obvious :=
  match goal with
    | [ |- ?P ] =>
      let t := tautReflect P in
        exact (tautTrue t)
  end.

(** We can verify that [obvious] solves our original example, with a proof term that does not mention details of the proof. *)
Adam Chlipala's avatar
Adam Chlipala committed
217
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
218 219 220 221 222 223 224

Theorem true_galore' : (True /\ True) -> (True \/ (True /\ (True -> True))).
  obvious.
Qed.

Print true_galore'.

Adam Chlipala's avatar
Adam Chlipala committed
225
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
226 227 228 229 230
true_galore' = 
tautTrue
  (TautImp (TautAnd TautTrue TautTrue)
     (TautOr TautTrue (TautAnd TautTrue (TautImp TautTrue TautTrue))))
     : True /\ True -> True \/ True /\ (True -> True)
Adam Chlipala's avatar
Adam Chlipala committed
231
 
Adam Chlipala's avatar
Adam Chlipala committed
232 233
    ]]

234
    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. *)
Adam Chlipala's avatar
Adam Chlipala committed
235 236


Adam Chlipala's avatar
Adam Chlipala committed
237 238
(** * A Monoid Expression Simplifier *)

239
(** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
240

Adam Chlipala's avatar
Adam Chlipala committed
241 242 243 244 245 246 247 248 249 250 251
Section monoid.
  Variable A : Set.
  Variable e : A.
  Variable f : A -> A -> A.

  Infix "+" := f.

  Hypothesis assoc : forall a b c, (a + b) + c = a + (b + c).
  Hypothesis identl : forall a, e + a = a.
  Hypothesis identr : forall a, a + e = a.

Adam Chlipala's avatar
Adam Chlipala committed
252 253
  (** 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.

254
     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. *)
Adam Chlipala's avatar
Adam Chlipala committed
255

Adam Chlipala's avatar
Adam Chlipala committed
256
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
257 258 259 260 261
  Inductive mexp : Set :=
  | Ident : mexp
  | Var : A -> mexp
  | Op : mexp -> mexp -> mexp.

262
  (** Next, we write an %``%#"#un-reflect#"#%''% function. *)
Adam Chlipala's avatar
Adam Chlipala committed
263

Adam Chlipala's avatar
Adam Chlipala committed
264 265 266 267 268 269 270
  Fixpoint mdenote (me : mexp) : A :=
    match me with
      | Ident => e
      | Var v => v
      | Op me1 me2 => mdenote me1 + mdenote me2
    end.

Adam Chlipala's avatar
Adam Chlipala committed
271 272
  (** 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. *)

Adam Chlipala's avatar
Adam Chlipala committed
273 274 275 276 277 278
  Fixpoint mldenote (ls : list A) : A :=
    match ls with
      | nil => e
      | x :: ls' => x + mldenote ls'
    end.

Adam Chlipala's avatar
Adam Chlipala committed
279 280
  (** The flattening function itself is easy to implement. *)

Adam Chlipala's avatar
Adam Chlipala committed
281 282 283 284 285 286 287
  Fixpoint flatten (me : mexp) : list A :=
    match me with
      | Ident => nil
      | Var x => x :: nil
      | Op me1 me2 => flatten me1 ++ flatten me2
    end.

Adam Chlipala's avatar
Adam Chlipala committed
288 289 290 291
  (** [flatten] has a straightforward correctness proof in terms of our [denote] functions. *)

  Lemma flatten_correct' : forall ml2 ml1,
    mldenote ml1 + mldenote ml2 = mldenote (ml1 ++ ml2).
Adam Chlipala's avatar
Adam Chlipala committed
292 293 294 295 296 297 298 299 300
    induction ml1; crush.
  Qed.

  Theorem flatten_correct : forall me, mdenote me = mldenote (flatten me).
    Hint Resolve flatten_correct'.

    induction me; crush.
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
301 302 303 304 305
  (** Now it is easy to prove a theorem that will be the main tool behind our simplification tactic. *)

  Theorem monoid_reflect : forall me1 me2,
    mldenote (flatten me1) = mldenote (flatten me2)
    -> mdenote me1 = mdenote me2.
Adam Chlipala's avatar
Adam Chlipala committed
306 307 308
    intros; repeat rewrite flatten_correct; assumption.
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
309 310 311 312
  (** We implement reflection into the [mexp] type. *)

  Ltac reflect me :=
    match me with
Adam Chlipala's avatar
Adam Chlipala committed
313
      | e => Ident
Adam Chlipala's avatar
Adam Chlipala committed
314 315 316
      | ?me1 + ?me2 =>
        let r1 := reflect me1 in
        let r2 := reflect me2 in
Adam Chlipala's avatar
Adam Chlipala committed
317
          constr:(Op r1 r2)
Adam Chlipala's avatar
Adam Chlipala committed
318
      | _ => constr:(Var me)
Adam Chlipala's avatar
Adam Chlipala committed
319 320
    end.

Adam Chlipala's avatar
Adam Chlipala committed
321 322
  (** The final [monoid] tactic works on goals that equate two monoid terms.  We reflect each and change the goal to refer to the reflected versions, finishing off by applying [monoid_reflect] and simplifying uses of [mldenote]. *)

Adam Chlipala's avatar
Adam Chlipala committed
323 324
  Ltac monoid :=
    match goal with
Adam Chlipala's avatar
Adam Chlipala committed
325 326 327
      | [ |- ?me1 = ?me2 ] =>
        let r1 := reflect me1 in
        let r2 := reflect me2 in
Adam Chlipala's avatar
Adam Chlipala committed
328 329 330 331
          change (mdenote r1 = mdenote r2);
            apply monoid_reflect; simpl mldenote
    end.

Adam Chlipala's avatar
Adam Chlipala committed
332 333
  (** We can make short work of theorems like this one: *)

Adam Chlipala's avatar
Adam Chlipala committed
334 335
(* end thide *)

Adam Chlipala's avatar
Adam Chlipala committed
336
  Theorem t1 : forall a b c d, a + b + c + d = a + (b + c) + d.
Adam Chlipala's avatar
Adam Chlipala committed
337 338 339 340
    intros; monoid.
    (** [[
  ============================
   a + (b + (c + (d + e))) = a + (b + (c + (d + e)))
Adam Chlipala's avatar
Adam Chlipala committed
341
 
Adam Chlipala's avatar
Adam Chlipala committed
342 343 344 345
        ]]

        [monoid] has canonicalized both sides of the equality, such that we can finish the proof by reflexivity. *)

Adam Chlipala's avatar
Adam Chlipala committed
346 347
    reflexivity.
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
348 349 350 351

  (** It is interesting to look at the form of the proof. *)

  Print t1.
Adam Chlipala's avatar
Adam Chlipala committed
352
  (** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
353 354 355 356 357 358
t1 = 
fun a b c d : A =>
monoid_reflect (Op (Op (Op (Var a) (Var b)) (Var c)) (Var d))
  (Op (Op (Var a) (Op (Var b) (Var c))) (Var d))
  (refl_equal (a + (b + (c + (d + e)))))
     : forall a b c d : A, a + b + c + d = a + (b + c) + d
Adam Chlipala's avatar
Adam Chlipala committed
359
 
Adam Chlipala's avatar
Adam Chlipala committed
360 361 362
      ]]

      The proof term contains only restatements of the equality operands in reflected form, followed by a use of reflexivity on the shared canonical form. *)
Adam Chlipala's avatar
Adam Chlipala committed
363

Adam Chlipala's avatar
Adam Chlipala committed
364 365
End monoid.

Adam Chlipala's avatar
Adam Chlipala committed
366 367
(** Extensions of this basic approach are used in the implementations of the [ring] and [field] tactics that come packaged with Coq. *)

Adam Chlipala's avatar
Adam Chlipala committed
368

Adam Chlipala's avatar
Adam Chlipala committed
369 370
(** * A Smarter Tautology Solver *)

Adam Chlipala's avatar
Adam Chlipala committed
371
(** 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.
Adam Chlipala's avatar
Adam Chlipala committed
372 373 374

   To arrive at a nice implementation satisfying these criteria, we introduce the [quote] tactic and its associated library. *)

Adam Chlipala's avatar
Adam Chlipala committed
375 376
Require Import Quote.

Adam Chlipala's avatar
Adam Chlipala committed
377
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
378 379 380 381 382 383 384 385
Inductive formula : Set :=
| Atomic : index -> formula
| Truth : formula
| Falsehood : formula
| And : formula -> formula -> formula
| Or : formula -> formula -> formula
| Imp : formula -> formula -> formula.

Adam Chlipala's avatar
Adam Chlipala committed
386 387 388
(** The type [index] comes from the [Quote] library and represents a countable variable type.  The rest of [formula]'s definition should be old hat by now.

   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. *)
Adam Chlipala's avatar
Adam Chlipala committed
389 390 391 392

Definition imp (P1 P2 : Prop) := P1 -> P2.
Infix "-->" := imp (no associativity, at level 95).

Adam Chlipala's avatar
Adam Chlipala committed
393 394 395 396
(** Now we can define our denotation function. *)

Definition asgn := varmap Prop.

Adam Chlipala's avatar
Adam Chlipala committed
397 398 399 400 401 402 403 404 405 406
Fixpoint formulaDenote (atomics : asgn) (f : formula) : Prop :=
  match f with
    | Atomic v => varmap_find False v atomics
    | Truth => True
    | Falsehood => False
    | And f1 f2 => formulaDenote atomics f1 /\ formulaDenote atomics f2
    | Or f1 f2 => formulaDenote atomics f1 \/ formulaDenote atomics f2
    | Imp f1 f2 => formulaDenote atomics f1 --> formulaDenote atomics f2
  end.

Adam Chlipala's avatar
Adam Chlipala committed
407 408
(** The [varmap] type family implements maps from [index] values.  In this case, we define an assignment as a map from variables to [Prop]s.  [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. *)

Adam Chlipala's avatar
Adam Chlipala committed
409 410 411 412 413
Section my_tauto.
  Variable atomics : asgn.

  Definition holds (v : index) := varmap_find False v atomics.

Adam Chlipala's avatar
Adam Chlipala committed
414 415
  (** 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. *)

Adam Chlipala's avatar
Adam Chlipala committed
416 417 418 419 420 421 422
  Require Import ListSet.

  Definition index_eq : forall x y : index, {x = y} + {x <> y}.
    decide equality.
  Defined.

  Definition add (s : set index) (v : index) := set_add index_eq v s.
Adam Chlipala's avatar
Adam Chlipala committed
423

Adam Chlipala's avatar
Adam Chlipala committed
424 425
  Definition In_dec : forall v (s : set index), {In v s} + {~ In v s}.
    Local Open Scope specif_scope.
Adam Chlipala's avatar
Adam Chlipala committed
426

Adam Chlipala's avatar
Adam Chlipala committed
427 428
    intro; refine (fix F (s : set index) : {In v s} + {~ In v s} :=
      match s with
Adam Chlipala's avatar
Adam Chlipala committed
429 430 431 432 433
        | nil => No
        | v' :: s' => index_eq v' v || F s'
      end); crush.
  Defined.

Adam Chlipala's avatar
Adam Chlipala committed
434 435
  (** We define what it means for all members of an index set to represent true propositions, and we prove some lemmas about this notion. *)

Adam Chlipala's avatar
Adam Chlipala committed
436 437 438 439 440 441 442 443 444 445 446 447 448 449 450 451 452 453 454 455 456 457 458 459 460
  Fixpoint allTrue (s : set index) : Prop :=
    match s with
      | nil => True
      | v :: s' => holds v /\ allTrue s'
    end.

  Theorem allTrue_add : forall v s,
    allTrue s
    -> holds v
    -> allTrue (add s v).
    induction s; crush;
      match goal with
        | [ |- context[if ?E then _ else _] ] => destruct E
      end; crush.
  Qed.

  Theorem allTrue_In : forall v s,
    allTrue s
    -> set_In v s
    -> varmap_find False v atomics.
    induction s; crush.
  Qed.

  Hint Resolve allTrue_add allTrue_In.

Adam Chlipala's avatar
Adam Chlipala committed
461
  Local Open Scope partial_scope.
Adam Chlipala's avatar
Adam Chlipala committed
462

Adam Chlipala's avatar
Adam Chlipala committed
463 464
  (** 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]. *)

Adam Chlipala's avatar
Adam Chlipala committed
465 466 467 468
  Definition forward (f : formula) (known : set index) (hyp : formula)
    (cont : forall known', [allTrue known' -> formulaDenote atomics f])
    : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f].
    refine (fix F (f : formula) (known : set index) (hyp : formula)
Adam Chlipala's avatar
Adam Chlipala committed
469
      (cont : forall known', [allTrue known' -> formulaDenote atomics f])
Adam Chlipala's avatar
Adam Chlipala committed
470
      : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f] :=
Adam Chlipala's avatar
Adam Chlipala committed
471
      match hyp with
Adam Chlipala's avatar
Adam Chlipala committed
472 473 474 475 476 477 478 479 480 481 482
        | Atomic v => Reduce (cont (add known v))
        | Truth => Reduce (cont known)
        | Falsehood => Yes
        | And h1 h2 =>
          Reduce (F (Imp h2 f) known h1 (fun known' =>
            Reduce (F f known' h2 cont)))
        | Or h1 h2 => F f known h1 cont && F f known h2 cont
        | Imp _ _ => Reduce (cont known)
      end); crush.
  Defined.

Adam Chlipala's avatar
Adam Chlipala committed
483 484
  (** A [backward] function implements analysis of the final goal.  It calls [forward] to handle implications. *)

Adam Chlipala's avatar
Adam Chlipala committed
485 486 487 488 489
  Definition backward (known : set index) (f : formula)
    : [allTrue known -> formulaDenote atomics f].
    refine (fix F (known : set index) (f : formula)
      : [allTrue known -> formulaDenote atomics f] :=
      match f with
Adam Chlipala's avatar
Adam Chlipala committed
490 491 492 493 494 495 496 497 498
        | Atomic v => Reduce (In_dec v known)
        | Truth => Yes
        | Falsehood => No
        | And f1 f2 => F known f1 && F known f2
        | Or f1 f2 => F known f1 || F known f2
        | Imp f1 f2 => forward f2 known f1 (fun known' => F known' f2)
      end); crush; eauto.
  Defined.

Adam Chlipala's avatar
Adam Chlipala committed
499 500
  (** A simple wrapper around [backward] gives us the usual type of a partial decision procedure. *)

Adam Chlipala's avatar
Adam Chlipala committed
501 502 503 504 505
  Definition my_tauto (f : formula) : [formulaDenote atomics f].
    intro; refine (Reduce (backward nil f)); crush.
  Defined.
End my_tauto.

Adam Chlipala's avatar
Adam Chlipala committed
506 507
(** 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 reflection for us.  Finally, we are able to construct an exact proof via [partialOut] and the [my_tauto] Gallina function. *)

Adam Chlipala's avatar
Adam Chlipala committed
508 509 510 511 512 513 514 515 516 517 518 519
Ltac my_tauto :=
  repeat match goal with
           | [ |- forall x : ?P, _ ] =>
             match type of P with
               | Prop => fail 1
               | _ => intro
             end
         end;
  quote formulaDenote;
  match goal with
    | [ |- formulaDenote ?m ?f ] => exact (partialOut (my_tauto m f))
  end.
Adam Chlipala's avatar
Adam Chlipala committed
520
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
521

Adam Chlipala's avatar
Adam Chlipala committed
522 523
(** A few examples demonstrate how the tactic works. *)

Adam Chlipala's avatar
Adam Chlipala committed
524 525 526 527 528
Theorem mt1 : True.
  my_tauto.
Qed.

Print mt1.
Adam Chlipala's avatar
Adam Chlipala committed
529
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
530 531
mt1 = partialOut (my_tauto (Empty_vm Prop) Truth)
     : True
Adam Chlipala's avatar
Adam Chlipala committed
532
 
Adam Chlipala's avatar
Adam Chlipala committed
533 534 535
    ]]

    We see [my_tauto] applied with an empty [varmap], since every subformula is handled by [formulaDenote]. *)
Adam Chlipala's avatar
Adam Chlipala committed
536 537 538 539 540 541

Theorem mt2 : forall x y : nat, x = y --> x = y.
  my_tauto.
Qed.

Print mt2.
Adam Chlipala's avatar
Adam Chlipala committed
542
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
543 544 545 546 547 548
mt2 = 
fun x y : nat =>
partialOut
  (my_tauto (Node_vm (x = y) (Empty_vm Prop) (Empty_vm Prop))
     (Imp (Atomic End_idx) (Atomic End_idx)))
     : forall x y : nat, x = y --> x = y
Adam Chlipala's avatar
Adam Chlipala committed
549
 
Adam Chlipala's avatar
Adam Chlipala committed
550 551 552
    ]]

    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. *)
Adam Chlipala's avatar
Adam Chlipala committed
553 554 555 556 557 558 559 560

Theorem mt3 : forall x y z,
  (x < y /\ y > z) \/ (y > z /\ x < S y)
  --> y > z /\ (x < y \/ x < S y).
  my_tauto.
Qed.

Print mt3.
Adam Chlipala's avatar
Adam Chlipala committed
561
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
562 563 564 565 566 567 568 569 570 571 572 573
fun x y z : nat =>
partialOut
  (my_tauto
     (Node_vm (x < S y) (Node_vm (x < y) (Empty_vm Prop) (Empty_vm Prop))
        (Node_vm (y > z) (Empty_vm Prop) (Empty_vm Prop)))
     (Imp
        (Or (And (Atomic (Left_idx End_idx)) (Atomic (Right_idx End_idx)))
           (And (Atomic (Right_idx End_idx)) (Atomic End_idx)))
        (And (Atomic (Right_idx End_idx))
           (Or (Atomic (Left_idx End_idx)) (Atomic End_idx)))))
     : forall x y z : nat,
       x < y /\ y > z \/ y > z /\ x < S y --> y > z /\ (x < y \/ x < S y)
Adam Chlipala's avatar
Adam Chlipala committed
574
 
Adam Chlipala's avatar
Adam Chlipala committed
575 576 577 578 579
    ]]

    Our goal contained three distinct atomic formulas, and we see that a three-element [varmap] is generated.

    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. *)
Adam Chlipala's avatar
Adam Chlipala committed
580 581 582 583 584 585

Theorem mt4 : True /\ True /\ True /\ True /\ True /\ True /\ False --> False.
  my_tauto.
Qed.

Print mt4.
Adam Chlipala's avatar
Adam Chlipala committed
586
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
587 588 589 590 591 592 593 594 595 596
mt4 = 
partialOut
  (my_tauto (Empty_vm Prop)
     (Imp
        (And Truth
           (And Truth
              (And Truth (And Truth (And Truth (And Truth Falsehood))))))
        Falsehood))
     : True /\ True /\ True /\ True /\ True /\ True /\ False --> False
    ]] *)
Adam Chlipala's avatar
Adam Chlipala committed
597 598 599 600 601 602

Theorem mt4' : True /\ True /\ True /\ True /\ True /\ True /\ False -> False.
  tauto.
Qed.

Print mt4'.
Adam Chlipala's avatar
Adam Chlipala committed
603
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
604 605 606 607 608 609 610 611 612 613 614 615 616 617 618 619 620
mt4' = 
fun H : True /\ True /\ True /\ True /\ True /\ True /\ False =>
and_ind
  (fun (_ : True) (H1 : True /\ True /\ True /\ True /\ True /\ False) =>
   and_ind
     (fun (_ : True) (H3 : True /\ True /\ True /\ True /\ False) =>
      and_ind
        (fun (_ : True) (H5 : True /\ True /\ True /\ False) =>
         and_ind
           (fun (_ : True) (H7 : True /\ True /\ False) =>
            and_ind
              (fun (_ : True) (H9 : True /\ False) =>
               and_ind (fun (_ : True) (H11 : False) => False_ind False H11)
                 H9) H7) H5) H3) H1) H
     : True /\ True /\ True /\ True /\ True /\ True /\ False -> False
    ]] *)

Adam Chlipala's avatar
Adam Chlipala committed
621 622 623

(** * Exercises *)

Adam Chlipala's avatar
Adam Chlipala committed
624 625
(** remove printing * *)

Adam Chlipala's avatar
Adam Chlipala committed
626 627 628 629
(** %\begin{enumerate}%#<ol>#

%\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.

Adam Chlipala's avatar
Adam Chlipala committed
630
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:
Adam Chlipala's avatar
Adam Chlipala committed
631 632 633 634 635 636 637

[[
  Theorem t2 : forall x y z, (2 # 1) * (x - (3 # 2) * y) == 15 # 1
    -> z + (8 # 1) * x == 20 # 1
    -> (-6 # 2) * y + (10 # 1) * x + z == 35 # 1.
    intros; reflectContext; assumption.
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
638
 
639 640
]]

Adam Chlipala's avatar
Adam Chlipala committed
641 642 643
  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 [==].

%\begin{enumerate}%#<ol>#
Adam Chlipala's avatar
Adam Chlipala committed
644
  %\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>#
Adam Chlipala's avatar
Adam Chlipala committed
645 646 647 648
  %\item%#<li># Define a function [lookup] for reading an element out of a list of rationals, by its position in the list.#</li>#
  %\item%#<li># Define a function [expDenote] that translates [exp]s, along with lists of rationals representing variable values, to [Q].#</li>#
  %\item%#<li># Define a recursive function [eqsDenote] over [list (exp * Q)], characterizing when all of the equations are true.#</li>#
  %\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>#
Adam Chlipala's avatar
Adam Chlipala committed
649
  %\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>#
Adam Chlipala's avatar
Adam Chlipala committed
650 651 652 653 654 655 656 657
  %\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>#
  %\item%#<li># Define a denotation function for [lhs].#</li>#
  %\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>#
  %\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>#
  %\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>#
  %\item%#<li># Write a tactic [reflect] to reflect a [Q] expression into [exp], with respect to a given list of variable values.#</li>#
  %\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>#
  %\item%#<li># Now this final tactic should do the job:
Adam Chlipala's avatar
Adam Chlipala committed
658

Adam Chlipala's avatar
Adam Chlipala committed
659 660 661 662 663 664 665 666 667 668 669 670 671 672 673 674 675 676 677 678 679 680 681 682
[[
  Ltac reflectContext :=
    let ls := findVarsHyps in
      repeat match goal with
               | [ H : ?e == ?num # ?den |- _ ] =>
                 let r := reflect ls e in
                   change (expDenote ls r == num # den) in H;
                   generalize H
             end;
      match goal with
        | [ |- ?g ] => let re := reflectEqs g in
            intros;
              let H := fresh "H" in
              assert (H : eqsDenote ls re); [ simpl in *; tauto
                | repeat match goal with
                           | [ H : expDenote _ _ == _ |- _ ] => clear H
                         end;
                generalize (linearizeEqsCorrect ls re H); clear H; simpl;
                  match goal with
                    | [ |- ?X == ?Y -> _ ] =>
                      ring_simplify X Y; intro
                  end ]
      end.

683 684
]]

Adam Chlipala's avatar
Adam Chlipala committed
685 686 687 688
#</ol>#%\end{enumerate}%
#</li>#
   
#</ol>#%\end{enumerate}% *)