Reflection.v 36.3 KB
Newer Older
1
(* Copyright (c) 2008-2012, 2015, Adam Chlipala
Adam Chlipala's avatar
Adam Chlipala committed
2 3 4 5 6 7 8 9 10 11 12
 * 
 * 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.

13
Require Import Cpdt.CpdtTactics Cpdt.MoreSpecif.
Adam Chlipala's avatar
Adam Chlipala committed
14 15

Set Implicit Arguments.
16
Set Asymmetric Patterns.
Adam Chlipala's avatar
Adam Chlipala committed
17 18 19 20 21
(* end hide *)


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

22
(** The last chapter highlighted a very heuristic approach to proving.  In this chapter, we will study an alternative technique,%\index{proof by reflection}% _proof by reflection_ %\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 _reflection_ applies because we will need to translate Gallina propositions into values of inductive types representing syntax, so that Gallina programs may analyze them, and translating such a term back to the original form is called _reflecting_ it. *)
Adam Chlipala's avatar
Adam Chlipala committed
23 24 25 26 27 28 29


(** * 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
30 31
| Even_O : isEven O
| Even_SS : forall n, isEven n -> isEven (S (S n)).
Adam Chlipala's avatar
Adam Chlipala committed
32

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

Theorem even_256 : isEven 256.
  prove_even.
Qed.

Print even_256.
Adam Chlipala's avatar
Adam Chlipala committed
42
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
43 44 45 46 47 48 49
even_256 = 
Even_SS
  (Even_SS
     (Even_SS
        (Even_SS
    ]]

50 51 52
    %\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.

    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.
Adam Chlipala's avatar
Adam Chlipala committed
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.

Adam Chlipala's avatar
Adam Chlipala committed
56
    The techniques of proof by reflection address both complaints.  We will be able to write proofs like in the example above 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
57

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

60
(* begin hide *)
61
(* begin thide *)
62
Definition paartial := partial.
63
(* end thide *)
64 65
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
66
Print partial.
Adam Chlipala's avatar
Adam Chlipala committed
67
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
68
Inductive partial (P : Prop) : Set :=  Proved : P -> [P] | Uncertain : [P]
Adam Chlipala's avatar
Adam Chlipala committed
69
    ]]
Adam Chlipala's avatar
Adam Chlipala committed
70

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

Adam Chlipala's avatar
Adam Chlipala committed
73
Local Open Scope partial_scope.
Adam Chlipala's avatar
Adam Chlipala committed
74 75 76

(** 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
77
(* begin thide *)
78
Definition check_even : forall n : nat, [isEven n].
Adam Chlipala's avatar
Adam Chlipala committed
79 80 81
  Hint Constructors isEven.

  refine (fix F (n : nat) : [isEven n] :=
Adam Chlipala's avatar
Adam Chlipala committed
82
    match n with
Adam Chlipala's avatar
Adam Chlipala committed
83 84 85 86 87 88
      | 0 => Yes
      | 1 => No
      | S (S n') => Reduce (F n')
    end); auto.
Defined.

89
(** The function [check_even] may be viewed as a _verified decision procedure_, because its type guarantees that it never returns %\coqdocnotation{%#<tt>#Yes#</tt>#%}% for inputs that are not even.
90 91

   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. *)
Adam Chlipala's avatar
Adam Chlipala committed
92 93 94 95 96 97 98 99 100 101

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

102
(** 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
103 104 105 106 107

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

110
(** 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]. *)
Adam Chlipala's avatar
Adam Chlipala committed
111 112 113 114 115 116

Theorem even_256' : isEven 256.
  prove_even_reflective.
Qed.

Print even_256'.
Adam Chlipala's avatar
Adam Chlipala committed
117
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
118 119 120 121
even_256' = partialOut (check_even 256)
     : isEven 256
    ]]

122 123 124
    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
125 126

Theorem even_255 : isEven 255.
127
  (** %\vspace{-.275in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
128
  prove_even_reflective.
129
]]
Adam Chlipala's avatar
Adam Chlipala committed
130

131
<<
Adam Chlipala's avatar
Adam Chlipala committed
132
User error: No matching clauses for match goal
133
>>
Adam Chlipala's avatar
Adam Chlipala committed
134 135 136

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

137
  %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
138
  exact (partialOut (check_even 255)).
139
]]
Adam Chlipala's avatar
Adam Chlipala committed
140

141
<<
Adam Chlipala's avatar
Adam Chlipala committed
142 143 144 145 146
  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"
147
>>
Adam Chlipala's avatar
Adam Chlipala committed
148

149
  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 %\coqdocnotation{%#<tt>#No#</tt>#%}%, so that the first term is equivalent to [True], which certainly does not unify with [isEven 255]. *)
Adam Chlipala's avatar
Adam Chlipala committed
150

Adam Chlipala's avatar
Adam Chlipala committed
151
Abort.
Adam Chlipala's avatar
Adam Chlipala committed
152

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

Adam Chlipala's avatar
Adam Chlipala committed
155

156
(** * Reifying the Syntax of a Trivial Tautology Language *)
Adam Chlipala's avatar
Adam Chlipala committed
157 158 159 160 161 162 163

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

164
(* begin hide *)
165
(* begin thide *)
166
Definition tg := (and_ind, or_introl).
167
(* end thide *)
168 169
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
170
Print true_galore.
Adam Chlipala's avatar
Adam Chlipala committed
171
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
172 173 174 175 176 177 178 179
true_galore = 
fun H : True /\ True =>
and_ind (fun _ _ : True => or_introl (True /\ (True -> True)) I) H
     : True /\ True -> True \/ True /\ (True -> True)
    ]]

    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.

180
   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}% _reify_ [Prop] into some type that we _can_ analyze.  This inductive type is a good candidate: *)
Adam Chlipala's avatar
Adam Chlipala committed
181

Adam Chlipala's avatar
Adam Chlipala committed
182
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
183 184 185 186 187 188
Inductive taut : Set :=
| TautTrue : taut
| TautAnd : taut -> taut -> taut
| TautOr : taut -> taut -> taut
| TautImp : taut -> taut -> taut.

189
(** We write a recursive function to _reflect_ this syntax back to [Prop].  Such functions are also called%\index{interpretation function}% _interpretation functions_, and we have used them in previous examples to give semantics to small programming languages. *)
Adam Chlipala's avatar
Adam Chlipala committed
190 191 192 193 194 195 196 197 198 199 200 201 202 203 204

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.

205
(** To use [tautTrue] to prove particular formulas, we need to implement the syntax reification process.  A recursive Ltac function does the job. *)
Adam Chlipala's avatar
Adam Chlipala committed
206

207
Ltac tautReify P :=
Adam Chlipala's avatar
Adam Chlipala committed
208 209 210
  match P with
    | True => TautTrue
    | ?P1 /\ ?P2 =>
211 212
      let t1 := tautReify P1 in
      let t2 := tautReify P2 in
Adam Chlipala's avatar
Adam Chlipala committed
213 214
        constr:(TautAnd t1 t2)
    | ?P1 \/ ?P2 =>
215 216
      let t1 := tautReify P1 in
      let t2 := tautReify P2 in
Adam Chlipala's avatar
Adam Chlipala committed
217 218
        constr:(TautOr t1 t2)
    | ?P1 -> ?P2 =>
219 220
      let t1 := tautReify P1 in
      let t2 := tautReify P2 in
Adam Chlipala's avatar
Adam Chlipala committed
221 222 223
        constr:(TautImp t1 t2)
  end.

224
(** With [tautReify] available, it is easy to finish our reflective tactic.  We look at the goal formula, reify it, and apply [tautTrue] to the reified formula. *)
Adam Chlipala's avatar
Adam Chlipala committed
225 226 227 228

Ltac obvious :=
  match goal with
    | [ |- ?P ] =>
229
      let t := tautReify P in
Adam Chlipala's avatar
Adam Chlipala committed
230 231 232 233
        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
234
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
235 236 237 238 239 240

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

Print true_galore'.
Adam Chlipala's avatar
Adam Chlipala committed
241
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
242 243 244 245 246 247 248
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
249
    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 _is_ 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 benefit is in addition to the proof-size improvement that we have already seen.
250 251

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


Adam Chlipala's avatar
Adam Chlipala committed
254 255
(** * A Monoid Expression Simplifier *)

256
(** 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
257

Adam Chlipala's avatar
Adam Chlipala committed
258 259 260 261 262 263 264 265 266 267 268
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
269 270
  (** 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.

271
     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
272

Adam Chlipala's avatar
Adam Chlipala committed
273
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
274 275 276 277 278
  Inductive mexp : Set :=
  | Ident : mexp
  | Var : A -> mexp
  | Op : mexp -> mexp -> mexp.

279
  (** Next, we write an interpretation function. *)
Adam Chlipala's avatar
Adam Chlipala committed
280

Adam Chlipala's avatar
Adam Chlipala committed
281 282 283 284 285 286 287
  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
288 289
  (** 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
290 291 292 293 294 295
  Fixpoint mldenote (ls : list A) : A :=
    match ls with
      | nil => e
      | x :: ls' => x + mldenote ls'
    end.

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

Adam Chlipala's avatar
Adam Chlipala committed
298 299 300 301 302 303 304
  Fixpoint flatten (me : mexp) : list A :=
    match me with
      | Ident => nil
      | Var x => x :: nil
      | Op me1 me2 => flatten me1 ++ flatten me2
    end.

305
  (** This function has a straightforward correctness proof in terms of our [denote] functions. *)
Adam Chlipala's avatar
Adam Chlipala committed
306 307 308

  Lemma flatten_correct' : forall ml2 ml1,
    mldenote ml1 + mldenote ml2 = mldenote (ml1 ++ ml2).
Adam Chlipala's avatar
Adam Chlipala committed
309 310 311 312 313 314 315 316 317
    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
318 319 320 321 322
  (** 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
323 324 325
    intros; repeat rewrite flatten_correct; assumption.
  Qed.

326
  (** We implement reification into the [mexp] type. *)
Adam Chlipala's avatar
Adam Chlipala committed
327

328
  Ltac reify me :=
Adam Chlipala's avatar
Adam Chlipala committed
329
    match me with
Adam Chlipala's avatar
Adam Chlipala committed
330
      | e => Ident
Adam Chlipala's avatar
Adam Chlipala committed
331
      | ?me1 + ?me2 =>
332 333
        let r1 := reify me1 in
        let r2 := reify me2 in
Adam Chlipala's avatar
Adam Chlipala committed
334
          constr:(Op r1 r2)
Adam Chlipala's avatar
Adam Chlipala committed
335
      | _ => constr:(Var me)
Adam Chlipala's avatar
Adam Chlipala committed
336 337
    end.

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

Adam Chlipala's avatar
Adam Chlipala committed
340 341
  Ltac monoid :=
    match goal with
Adam Chlipala's avatar
Adam Chlipala committed
342
      | [ |- ?me1 = ?me2 ] =>
343 344
        let r1 := reify me1 in
        let r2 := reify me2 in
Adam Chlipala's avatar
Adam Chlipala committed
345
          change (mdenote r1 = mdenote r2);
346
            apply monoid_reflect; simpl
Adam Chlipala's avatar
Adam Chlipala committed
347 348
    end.

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

Adam Chlipala's avatar
Adam Chlipala committed
351 352
(* end thide *)

Adam Chlipala's avatar
Adam Chlipala committed
353
  Theorem t1 : forall a b c d, a + b + c + d = a + (b + c) + d.
Adam Chlipala's avatar
Adam Chlipala committed
354 355 356 357
    intros; monoid.
    (** [[
  ============================
   a + (b + (c + (d + e))) = a + (b + (c + (d + e)))
Adam Chlipala's avatar
Adam Chlipala committed
358
 
Adam Chlipala's avatar
Adam Chlipala committed
359 360
        ]]

361
        Our tactic has canonicalized both sides of the equality, such that we can finish the proof by reflexivity. *)
Adam Chlipala's avatar
Adam Chlipala committed
362

Adam Chlipala's avatar
Adam Chlipala committed
363 364
    reflexivity.
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
365 366 367 368

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

  Print t1.
Adam Chlipala's avatar
Adam Chlipala committed
369
  (** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
370 371 372 373
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))
374
  (eq_refl (a + (b + (c + (d + e)))))
Adam Chlipala's avatar
Adam Chlipala committed
375 376 377
     : forall a b c d : A, a + b + c + d = a + (b + c) + d
      ]]

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

Adam Chlipala's avatar
Adam Chlipala committed
380 381
End monoid.

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

Adam Chlipala's avatar
Adam Chlipala committed
384

Adam Chlipala's avatar
Adam Chlipala committed
385 386
(** * A Smarter Tautology Solver *)

387
(** 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
388

389
   To arrive at a nice implementation satisfying these criteria, we introduce the %\index{tactics!quote}%[quote] tactic and its associated library. *)
Adam Chlipala's avatar
Adam Chlipala committed
390

Adam Chlipala's avatar
Adam Chlipala committed
391 392
Require Import Quote.

Adam Chlipala's avatar
Adam Chlipala committed
393
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
394 395 396 397 398 399 400
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
401
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
402

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

405
   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 wants to treat function types specially, so it gets confused if function types are part of the structure we want to encode syntactically.  To trick [quote] into not noticing our uses of function types to express logical implication, we will need to declare a wrapper definition for implication, as we did in the last chapter. *)
Adam Chlipala's avatar
Adam Chlipala committed
406 407 408 409

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

Adam Chlipala's avatar
Adam Chlipala committed
410 411 412 413
(** Now we can define our denotation function. *)

Definition asgn := varmap Prop.

Adam Chlipala's avatar
Adam Chlipala committed
414
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
415 416 417 418 419 420 421 422 423
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
424
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
425

426
(** 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 interpretation function [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
427

Adam Chlipala's avatar
Adam Chlipala committed
428 429 430 431 432
Section my_tauto.
  Variable atomics : asgn.

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

Adam Chlipala's avatar
Adam Chlipala committed
433 434
  (** 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
435 436 437 438 439 440 441
  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
442

Adam Chlipala's avatar
Adam Chlipala committed
443 444
  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
445

Adam Chlipala's avatar
Adam Chlipala committed
446 447
    intro; refine (fix F (s : set index) : {In v s} + {~ In v s} :=
      match s with
Adam Chlipala's avatar
Adam Chlipala committed
448 449 450 451 452
        | nil => No
        | v' :: s' => index_eq v' v || F s'
      end); crush.
  Defined.

Adam Chlipala's avatar
Adam Chlipala committed
453 454
  (** 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
455 456 457 458 459 460 461 462 463 464 465 466 467 468 469 470 471 472 473 474 475 476 477 478 479
  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
480
  Local Open Scope partial_scope.
Adam Chlipala's avatar
Adam Chlipala committed
481

482 483 484
  (** Now we can write a function [forward] that implements deconstruction of hypotheses, expanding a compound formula into a set of sets of atomic formulas covering all possible cases introduced with use of [Or].  To handle consideration of multiple cases, the function takes in a continuation argument, which will be called once for each case.

     The [forward] function 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
485

486 487 488
  Definition forward : forall (f : formula) (known : set index) (hyp : formula)
    (cont : forall known', [allTrue known' -> formulaDenote atomics f]),
    [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f].
Adam Chlipala's avatar
Adam Chlipala committed
489
    refine (fix F (f : formula) (known : set index) (hyp : formula)
Adam Chlipala's avatar
Adam Chlipala committed
490
      (cont : forall known', [allTrue known' -> formulaDenote atomics f])
Adam Chlipala's avatar
Adam Chlipala committed
491
      : [allTrue known -> formulaDenote atomics hyp -> formulaDenote atomics f] :=
Adam Chlipala's avatar
Adam Chlipala committed
492
      match hyp with
Adam Chlipala's avatar
Adam Chlipala committed
493 494 495 496 497 498 499 500 501 502 503
        | 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
504 505
  (** A [backward] function implements analysis of the final goal.  It calls [forward] to handle implications. *)

Adam Chlipala's avatar
Adam Chlipala committed
506
(* begin thide *)
507 508
  Definition backward : forall (known : set index) (f : formula),
    [allTrue known -> formulaDenote atomics f].
Adam Chlipala's avatar
Adam Chlipala committed
509 510 511
    refine (fix F (known : set index) (f : formula)
      : [allTrue known -> formulaDenote atomics f] :=
      match f with
Adam Chlipala's avatar
Adam Chlipala committed
512 513 514 515 516 517 518 519
        | 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
520
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
521

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

524
  Definition my_tauto : forall f : formula, [formulaDenote atomics f].
Adam Chlipala's avatar
Adam Chlipala committed
525
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
526 527
    intro; refine (Reduce (backward nil f)); crush.
  Defined.
Adam Chlipala's avatar
Adam Chlipala committed
528
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
529 530
End my_tauto.

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

Adam Chlipala's avatar
Adam Chlipala committed
533 534 535 536 537 538 539 540 541 542 543 544
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
545
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
546

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

Adam Chlipala's avatar
Adam Chlipala committed
549 550 551 552 553
Theorem mt1 : True.
  my_tauto.
Qed.

Print mt1.
Adam Chlipala's avatar
Adam Chlipala committed
554
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
555 556 557 558 559
mt1 = partialOut (my_tauto (Empty_vm Prop) Truth)
     : True
    ]]

    We see [my_tauto] applied with an empty [varmap], since every subformula is handled by [formulaDenote]. *)
Adam Chlipala's avatar
Adam Chlipala committed
560 561 562 563 564

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

565
(* begin hide *)
566
(* begin thide *)
567
Definition nvm := (Node_vm, Empty_vm, End_idx, Left_idx, Right_idx).
568
(* end thide *)
569 570
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
571
Print mt2.
Adam Chlipala's avatar
Adam Chlipala committed
572
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
573 574 575 576 577 578 579 580 581
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
    ]]

    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
582 583 584 585 586 587 588 589

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
590
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
591 592 593 594 595 596 597 598 599 600 601 602 603 604 605 606 607
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)
    ]]

    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
608 609 610 611 612 613

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

Print mt4.
Adam Chlipala's avatar
Adam Chlipala committed
614
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
615 616 617 618 619 620 621 622 623
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
624 625
    ]]
    *)
Adam Chlipala's avatar
Adam Chlipala committed
626 627 628 629 630

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

631
(* begin hide *)
632
(* begin thide *)
633
Definition fi := False_ind.
634
(* end thide *)
635 636
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
637
Print mt4'.
Adam Chlipala's avatar
Adam Chlipala committed
638
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
639 640 641 642 643 644 645 646 647 648 649 650 651 652 653
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
654
    ]]
655 656

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

658 659
(** ** Manual Reification of Terms with Variables *)

Adam Chlipala's avatar
Adam Chlipala committed
660
(* begin thide *)
661 662
(** The action of the [quote] tactic above may seem like magic.  Somehow it performs equality comparison between subterms of arbitrary types, so that these subterms may be represented with the same reified variable.  While [quote] is implemented in OCaml, we can code the reification process completely in Ltac, as well.  To make our job simpler, we will represent variables as [nat]s, indexing into a simple list of variable values that may be referenced.

663
   Step one of the process is to crawl over a term, building a duplicate-free list of all values that appear in positions we will encode as variables.  A useful helper function adds an element to a list, preventing duplicates.  Note how we use Ltac pattern matching to implement an equality test on Gallina terms; this is simple syntactic equality, not even the richer definitional equality.  We also represent lists as nested tuples, to allow different list elements to have different Gallina types. *)
664 665 666 667 668 669 670 671 672 673 674 675

Ltac inList x xs :=
  match xs with
    | tt => false
    | (x, _) => true
    | (_, ?xs') => inList x xs'
  end.

Ltac addToList x xs :=
  let b := inList x xs in
    match b with
      | true => xs
Adam Chlipala's avatar
Adam Chlipala committed
676
      | false => constr:((x, xs))
677 678 679 680 681 682 683 684 685 686 687 688 689 690 691 692 693 694 695 696 697 698 699 700 701 702 703 704 705 706
    end.

(** Now we can write our recursive function to calculate the list of variable values we will want to use to represent a term. *)

Ltac allVars xs e :=
  match e with
    | True => xs
    | False => xs
    | ?e1 /\ ?e2 =>
      let xs := allVars xs e1 in
        allVars xs e2
    | ?e1 \/ ?e2 =>
      let xs := allVars xs e1 in
        allVars xs e2
    | ?e1 -> ?e2 =>
      let xs := allVars xs e1 in
        allVars xs e2
    | _ => addToList e xs
  end.

(** We will also need a way to map a value to its position in a list. *)

Ltac lookup x xs :=
  match xs with
    | (x, _) => O
    | (_, ?xs') =>
      let n := lookup x xs' in
        constr:(S n)
  end.

707
(** The next building block is a procedure for reifying a term, given a list of all allowed variable values.  We are free to make this procedure partial, where tactic failure may be triggered upon attempting to reify a term containing subterms not included in the list of variables.  The type of the output term is a copy of [formula] where [index] is replaced by [nat], in the type of the constructor for atomic formulas. *)
708 709 710 711 712 713 714 715 716 717 718 719 720

Inductive formula' : Set :=
| Atomic' : nat -> formula'
| Truth' : formula'
| Falsehood' : formula'
| And' : formula' -> formula' -> formula'
| Or' : formula' -> formula' -> formula'
| Imp' : formula' -> formula' -> formula'.

(** Note that, when we write our own Ltac procedure, we can work directly with the normal [->] operator, rather than needing to introduce a wrapper for it. *)

Ltac reifyTerm xs e :=
  match e with
Adam Chlipala's avatar
Adam Chlipala committed
721 722
    | True => Truth'
    | False => Falsehood'
723 724 725 726 727 728 729 730 731 732 733 734 735 736 737 738 739 740 741 742 743 744 745 746 747 748 749 750 751 752 753 754 755 756 757 758 759 760 761 762 763 764 765
    | ?e1 /\ ?e2 =>
      let p1 := reifyTerm xs e1 in
      let p2 := reifyTerm xs e2 in
        constr:(And' p1 p2)
    | ?e1 \/ ?e2 =>
      let p1 := reifyTerm xs e1 in
      let p2 := reifyTerm xs e2 in
        constr:(Or' p1 p2)
    | ?e1 -> ?e2 =>
      let p1 := reifyTerm xs e1 in
      let p2 := reifyTerm xs e2 in
        constr:(Imp' p1 p2)
    | _ =>
      let n := lookup e xs in
        constr:(Atomic' n)
  end.

(** Finally, we bring all the pieces together. *)

Ltac reify :=
  match goal with
    | [ |- ?G ] => let xs := allVars tt G in
      let p := reifyTerm xs G in
        pose p
  end.

(** A quick test verifies that we are doing reification correctly. *)

Theorem mt3' : forall x y z,
  (x < y /\ y > z) \/ (y > z /\ x < S y)
  -> y > z /\ (x < y \/ x < S y).
  do 3 intro; reify.

(** Our simple tactic adds the translated term as a new variable:
[[
f := Imp'
         (Or' (And' (Atomic' 2) (Atomic' 1)) (And' (Atomic' 1) (Atomic' 0)))
         (And' (Atomic' 1) (Or' (Atomic' 2) (Atomic' 0))) : formula'
]]
*)
Abort.

(** More work would be needed to complete the reflective tactic, as we must connect our new syntax type with the real meanings of formulas, but the details are the same as in our prior implementation with [quote]. *)
Adam Chlipala's avatar
Adam Chlipala committed
766
(* end thide *)
767 768 769 770 771 772 773 774 775 776 777 778 779 780 781 782 783 784 785 786 787 788 789 790 791 792 793 794


(** * Building a Reification Tactic that Recurses Under Binders *)

(** All of our examples so far have stayed away from reifying the syntax of terms that use such features as quantifiers and [fun] function abstractions.  Such cases are complicated by the fact that different subterms may be allowed to reference different sets of free variables.  Some cleverness is needed to clear this hurdle, but a few simple patterns will suffice.  Consider this example of a simple dependently typed term language, where a function abstraction body is represented conveniently with a Coq function. *)

Inductive type : Type :=
| Nat : type
| NatFunc : type -> type.

Inductive term : type -> Type :=
| Const : nat -> term Nat
| Plus : term Nat -> term Nat -> term Nat
| Abs : forall t, (nat -> term t) -> term (NatFunc t).

Fixpoint typeDenote (t : type) : Type :=
  match t with
    | Nat => nat
    | NatFunc t => nat -> typeDenote t
  end.

Fixpoint termDenote t (e : term t) : typeDenote t :=
  match e with
    | Const n => n
    | Plus e1 e2 => termDenote e1 + termDenote e2
    | Abs _ e1 => fun x => termDenote (e1 x)
  end.

795
(** Here is a %\%naive%{}% first attempt at a reification tactic. *)
796

797 798 799
(* begin hide *)
Definition red_herring := O.
(* end hide *)
800 801 802 803 804 805 806 807 808 809 810 811 812 813
Ltac refl' e :=
  match e with
    | ?E1 + ?E2 =>
      let r1 := refl' E1 in
      let r2 := refl' E2 in
        constr:(Plus r1 r2)

    | fun x : nat => ?E1 =>
      let r1 := refl' E1 in
        constr:(Abs (fun x => r1 x))

    | _ => constr:(Const e)
  end.

814
(** Recall that a regular Ltac pattern variable [?X] only matches terms that _do not mention new variables introduced within the pattern_.  In our %\%naive%{}% implementation, the case for matching function abstractions matches the function body in a way that prevents it from mentioning the function argument!  Our code above plays fast and loose with the function body in a way that leads to independent problems, but we could change the code so that it indeed handles function abstractions that ignore their arguments.
815

816
   To handle functions in general, we will use the pattern variable form [@?X], which allows [X] to mention newly introduced variables that are declared explicitly.  A use of [@?X] must be followed by a list of the local variables that may be mentioned.  The variable [X] then comes to stand for a Gallina function over the values of those variables.  For instance: *)
817 818

Reset refl'.
819 820 821 822
(* begin hide *)
Reset red_herring.
Definition red_herring := O.
(* end hide *)
823 824 825 826 827 828 829 830 831 832 833 834 835 836
Ltac refl' e :=
  match e with
    | ?E1 + ?E2 =>
      let r1 := refl' E1 in
      let r2 := refl' E2 in
        constr:(Plus r1 r2)

    | fun x : nat => @?E1 x =>
      let r1 := refl' E1 in
        constr:(Abs r1)

    | _ => constr:(Const e)
  end.

837
(** Now, in the abstraction case, we bind [E1] as a function from an [x] value to the value of the abstraction body.  Unfortunately, our recursive call there is not destined for success.  It will match the same abstraction pattern and trigger another recursive call, and so on through infinite recursion.  One last refactoring yields a working procedure.  The key idea is to consider every input to [refl'] as _a function over the values of variables introduced during recursion_. *)
838 839

Reset refl'.
840 841 842
(* begin hide *)
Reset red_herring.
(* end hide *)
843 844 845 846 847 848 849 850 851
Ltac refl' e :=
  match eval simpl in e with
    | fun x : ?T => @?E1 x + @?E2 x =>
      let r1 := refl' E1 in
      let r2 := refl' E2 in
        constr:(fun x => Plus (r1 x) (r2 x))

    | fun (x : ?T) (y : nat) => @?E1 x y =>
      let r1 := refl' (fun p : T * nat => E1 (fst p) (snd p)) in
852
        constr:(fun u => Abs (fun v => r1 (u, v)))
853 854 855 856 857 858 859 860 861 862 863 864 865 866 867 868 869 870 871 872 873 874 875 876 877 878 879 880 881 882 883

    | _ => constr:(fun x => Const (e x))
  end.

(** Note how now even the addition case works in terms of functions, with [@?X] patterns.  The abstraction case introduces a new variable by extending the type used to represent the free variables.  In particular, the argument to [refl'] used type [T] to represent all free variables.  We extend the type to [T * nat] for the type representing free variable values within the abstraction body.  A bit of bookkeeping with pairs and their projections produces an appropriate version of the abstraction body to pass in a recursive call.  To ensure that all this repackaging of terms does not interfere with pattern matching, we add an extra [simpl] reduction on the function argument, in the first line of the body of [refl'].

   Now one more tactic provides an example of how to apply reification.  Let us consider goals that are equalities between terms that can be reified.  We want to change such goals into equalities between appropriate calls to [termDenote]. *)

Ltac refl :=
  match goal with
    | [ |- ?E1 = ?E2 ] =>
      let E1' := refl' (fun _ : unit => E1) in
      let E2' := refl' (fun _ : unit => E2) in
        change (termDenote (E1' tt) = termDenote (E2' tt));
          cbv beta iota delta [fst snd]
  end.

Goal (fun (x y : nat) => x + y + 13) = (fun (_ z : nat) => z).
  refl.
(** %\vspace{-.15in}%[[
  ============================
   termDenote
     (Abs
        (fun y : nat =>
         Abs (fun y0 : nat => Plus (Plus (Const y) (Const y0)) (Const 13)))) =
   termDenote (Abs (fun _ : nat => Abs (fun y0 : nat => Const y0)))
]]
*)

Abort.

884
(** Our encoding here uses Coq functions to represent binding within the terms we reify, which makes it difficult to implement certain functions over reified terms.  An alternative would be to represent variables with numbers.  This can be done by writing a slightly smarter reification function that identifies variable references by detecting when term arguments are just compositions of [fst] and [snd]; from the order of the compositions we may read off the variable number.  We leave the details as an exercise (though not a trivial one!) for the reader. *)