add closedNDtoNormalND to NaturalDeduction
[coq-hetmet.git] / src / NaturalDeduction.v
1 (*********************************************************************************************************************************)
2 (* NaturalDeduction: structurally explicit proofs in Coq                                                                         *)
3 (*********************************************************************************************************************************)
4
5 Generalizable All Variables.
6 Require Import Preamble.
7 Require Import General.
8 Require Import Coq.Strings.Ascii.
9 Require Import Coq.Strings.String.
10
11 (*
12  * IMPORTANT!!!
13  *
14  * Unlike most formalizations, this library offers TWO different ways
15  * to represent a natural deduction proof.  To demonstrate this,
16  * consider the signature of the propositional calculus:
17  *
18  *   Variable  PropositionalVariable : Type.
19  *
20  *   Inductive Formula : Prop :=
21  *   | formula_var : PropositionalVariable -> Formula   (* every propositional variable is a formula *)
22  *   | formula_and :   Formula ->  Formula -> Formula   (* the conjunction of any two formulae is a formula *)
23  *   | formula_or  :   Formula ->  Formula -> Formula   (* the disjunction of any two formulae is a formula *)
24  *
25  * And couple this with the theory of conjunction and disjunction:
26  * φ\/ψ is true if either φ is true or ψ is true, and φ/\ψ is true
27  * if both φ and ψ are true.
28  *
29  * 1) Structurally implicit proofs
30  *
31  *    This is what you would call the "usual" representation –- it's
32  *    what most people learn when they first start programming in Coq:
33  *
34  *    Inductive IsTrue : Formula -> Prop :=
35  *    | IsTrue_or1 : forall f1 f2, IsTrue f1 ->              IsTrue (formula_or  f1 f2) 
36  *    | IsTrue_or2 : forall f1 f2,              IsTrue f2 -> IsTrue (formula_or  f1 f2) 
37  *    | IsTrue_and : forall f1 f2, IsTrue f2 -> IsTrue f2 -> IsTrue (formula_and f1 f2) 
38  *
39  *    Here each judgment (such as "φ is true") is represented by a Coq
40  *    type; furthermore:
41  *
42  *       1. A proof of a judgment is any inhabitant of that Coq type.
43  *
44  *       2. A proof of a judgment "J2" from hypothesis judgment "J1"
45  *          is any Coq function from the Coq type for J1 to the Coq
46  *          type for J2; Composition of (hypothetical) proofs is
47  *          represented by composition of Coq functions.
48  *
49  *       3. A pair of judgments is represented by their product (Coq
50  *          type [prod]) in Coq; a pair of proofs is represented by
51  *          their pair (Coq function [pair]) in Coq.
52  *
53  *       4. Duplication of hypotheses is represented by the Coq
54  *          function (fun x => (x,x)).  Dereliction of hypotheses is
55  *          represented by the coq function (fun (x,y) => x) or (fun
56  *          (x,y) => y).  Exchange of the order of hypotheses is
57  *          represented by the Coq function (fun (x,y) => (y,x)).
58  *
59  *    The above can be done using lists instead of tuples.
60  *
61  *    The advantage of this approach is that it requires a minimum
62  *    amount of syntax, and takes maximum advantage of Coq's
63  *    automation facilities.
64  *
65  *    The disadvantage is that one cannot reason about proof-theoretic
66  *    properties *generically* across different signatures and
67  *    theories.  Each signature has its own type of judgments, and
68  *    each theory has its own type of proofs.  In the present
69  *    development we will want to prove –– in this generic manner --
70  *    that various classes of natural deduction calculi form various
71  *    kinds of categories.  So we will need this ability to reason
72  *    about proofs independently of the type used to represent
73  *    judgments and (more importantly) of the set of basic inference
74  *    rules.
75  *
76  * 2) Structurally explicit proofs
77  *
78  *    Structurally explicit proofs are formalized in this file
79  *    (NaturalDeduction.v) and are designed specifically in order to
80  *    circumvent the problem in the previous paragraph.
81  *
82  *)
83
84 (*
85  * REGARDING LISTS versus TREES:
86  *
87  * You'll notice that this formalization uses (Tree (option A)) in a
88  * lot of places where you might find (list A) more natural.  If this
89  * bothers you, see the end of the file for the technical reasons why.
90  * In short, it lets us avoid having to mess about with JMEq or EqDep,
91  * which are not as well-supported by the Coq core as the theory of
92  * CiC proper.
93  *)
94
95 Section Natural_Deduction.
96
97   (* any Coq Type may be used as the set of judgments *)
98   Context {Judgment : Type}.
99
100   (* any Coq Type –- indexed by the hypothesis and conclusion judgments -- may be used as the set of basic inference rules *)
101   Context {Rule     : forall (hypotheses:Tree ??Judgment)(conclusion:Tree ??Judgment), Type}.
102
103   (*
104    *  This type represents a valid Natural Deduction proof from a list
105    *  (tree) of hypotheses; the notation H/⋯⋯/C is meant to look like
106    *  a proof tree with the middle missing if you tilt your head to
107    *  the left (yeah, I know it's a stretch).  Note also that this
108    *  type is capable of representing proofs with multiple
109    *  conclusions, whereas a Rule may have only one conclusion.
110    *) 
111   Inductive ND :
112     forall hypotheses:Tree ??Judgment,
113       forall conclusions:Tree ??Judgment,
114         Type :=
115
116     (* natural deduction: you may infer anything from itself -- "identity proof" *)
117     | nd_id0    :             [   ] /⋯⋯/ [   ]
118     | nd_id1    : forall  h,  [ h ] /⋯⋯/ [ h ]
119   
120     (* natural deduction: you may discard conclusions *)
121     | nd_weak   : forall  h,  [ h ] /⋯⋯/ [   ]
122   
123     (* natural deduction: you may duplicate conclusions *)
124     | nd_copy   : forall  h,    h   /⋯⋯/ (h,,h)
125   
126     (* natural deduction: you may write two proof trees side by side on a piece of paper -- "proof product" *)
127     | nd_prod : forall {h1 h2 c1 c2}
128        (pf1: h1       /⋯⋯/ c1      )
129        (pf2:       h2 /⋯⋯/       c2),
130        (     h1 ,, h2 /⋯⋯/ c1 ,, c2)
131   
132     (* natural deduction: given a proof of every hypothesis, you may discharge them -- "proof composition" *)
133     | nd_comp :
134       forall {h x c}
135       `(pf1: h /⋯⋯/ x)
136       `(pf2: x /⋯⋯/ c),
137        (     h /⋯⋯/ c)
138   
139     (* structural rules on lists of judgments *)
140     | nd_cancell : forall {a},       [] ,, a /⋯⋯/ a
141     | nd_cancelr : forall {a},       a ,, [] /⋯⋯/ a
142     | nd_llecnac : forall {a},             a /⋯⋯/ [] ,, a
143     | nd_rlecnac : forall {a},             a /⋯⋯/ a ,, []
144     | nd_assoc   : forall {a b c}, (a,,b),,c /⋯⋯/ a,,(b,,c)
145     | nd_cossa   : forall {a b c}, a,,(b,,c) /⋯⋯/ (a,,b),,c
146
147     (* any Rule by itself counts as a proof *)
148     | nd_rule    : forall {h c} (r:Rule h c), h /⋯⋯/ c
149   
150     where  "H /⋯⋯/ C" := (ND H C).
151
152     Notation "H /⋯⋯/ C" := (ND H C) : pf_scope.
153     Notation "a ;; b"   := (nd_comp a b) : nd_scope.
154     Notation "a ** b"   := (nd_prod a b) : nd_scope.
155     Open Scope nd_scope.
156     Open Scope pf_scope.
157
158   (* a proof is "structural" iff it does not contain any invocations of nd_rule *)
159   Inductive Structural : forall {h c}, h /⋯⋯/ c -> Prop :=
160   | nd_structural_id0     :                                                                            Structural nd_id0
161   | nd_structural_id1     : forall h,                                                                  Structural (nd_id1 h)
162   | nd_structural_weak    : forall h,                                                                  Structural (nd_weak h)
163   | nd_structural_copy    : forall h,                                                                  Structural (nd_copy h)
164   | nd_structural_prod    : forall `(pf1:h1/⋯⋯/c1)`(pf2:h2/⋯⋯/c2), Structural pf1 -> Structural pf2 -> Structural (pf1**pf2)
165   | nd_structural_comp    : forall `(pf1:h1/⋯⋯/x) `(pf2: x/⋯⋯/c2), Structural pf1 -> Structural pf2 -> Structural (pf1;;pf2)
166   | nd_structural_cancell : forall {a},                                                                Structural (@nd_cancell a)
167   | nd_structural_cancelr : forall {a},                                                                Structural (@nd_cancelr a)
168   | nd_structural_llecnac : forall {a},                                                                Structural (@nd_llecnac a)
169   | nd_structural_rlecnac : forall {a},                                                                Structural (@nd_rlecnac a)
170   | nd_structural_assoc   : forall {a b c},                                                            Structural (@nd_assoc a b c)
171   | nd_structural_cossa   : forall {a b c},                                                            Structural (@nd_cossa a b c)
172   .
173
174   (* multi-judgment generalization of nd_id0 and nd_id1; making nd_id0/nd_id1 primitive and deriving all other *)
175   Fixpoint nd_id (sl:Tree ??Judgment) : sl /⋯⋯/ sl :=
176     match sl with
177       | T_Leaf None      => nd_id0
178       | T_Leaf (Some x)  => nd_id1 x
179       | T_Branch a b     => nd_prod (nd_id a) (nd_id b)
180     end.
181
182   Hint Constructors Structural.
183   Lemma nd_id_structural : forall sl, Structural (nd_id sl).
184     intros.
185     induction sl; simpl; auto.
186     destruct a; auto.
187     Defined.
188
189   (* An equivalence relation on proofs which is sensitive only to the logical content of the proof -- insensitive to
190    * structural variations  *)
191   Class ND_Relation :=
192   { ndr_eqv                  : forall {h c  }, h /⋯⋯/ c -> h /⋯⋯/ c -> Prop where "pf1 === pf2" := (@ndr_eqv _ _  pf1 pf2)
193   ; ndr_eqv_equivalence      : forall h c, Equivalence (@ndr_eqv h c)
194
195   (* the relation must respect composition, be associative wrt composition, and be left and right neutral wrt the identity proof *)
196   ; ndr_comp_respects        : forall {a b c}(f f':a/⋯⋯/b)(g g':b/⋯⋯/c),      f === f' -> g === g' -> f;;g === f';;g'
197   ; ndr_comp_associativity   : forall `(f:a/⋯⋯/b)`(g:b/⋯⋯/c)`(h:c/⋯⋯/d),                         (f;;g);;h === f;;(g;;h)
198   ; ndr_comp_left_identity   : forall `(f:a/⋯⋯/c),                                          nd_id _ ;; f   === f
199   ; ndr_comp_right_identity  : forall `(f:a/⋯⋯/c),                                          f ;; nd_id _   === f
200
201   (* the relation must respect products, be associative wrt products, and be left and right neutral wrt the identity proof *)
202   ; ndr_prod_respects        : forall {a b c d}(f f':a/⋯⋯/b)(g g':c/⋯⋯/d),     f===f' -> g===g' ->    f**g === f'**g'
203   ; ndr_prod_associativity   : forall `(f:a/⋯⋯/a')`(g:b/⋯⋯/b')`(h:c/⋯⋯/c'),       (f**g)**h === nd_assoc ;; f**(g**h) ;; nd_cossa
204   ; ndr_prod_left_identity   : forall `(f:a/⋯⋯/b),                       (nd_id0 ** f ) === nd_cancell ;; f ;; nd_llecnac
205   ; ndr_prod_right_identity  : forall `(f:a/⋯⋯/b),                       (f ** nd_id0)  === nd_cancelr ;; f ;; nd_rlecnac
206
207   (* products and composition must distribute over each other *)
208   ; ndr_prod_preserves_comp  : forall `(f:a/⋯⋯/b)`(f':a'/⋯⋯/b')`(g:b/⋯⋯/c)`(g':b'/⋯⋯/c'), (f;;g)**(f';;g') === (f**f');;(g**g')
209
210   (* any two _structural_ proofs with the same hypotheses/conclusions must be considered equal *)
211   ; ndr_structural_indistinguishable : forall `(f:a/⋯⋯/b)(g:a/⋯⋯/b), Structural f -> Structural g -> f===g
212   }.
213
214   (* 
215    * Single-conclusion proofs; this is an alternate representation
216    * where each inference has only a single conclusion.  These have
217    * worse compositionality properties than ND's, but are easier to
218    * emit as LaTeX code.
219    *)
220   Inductive SCND : Tree ??Judgment -> Tree ??Judgment -> Type :=
221   | scnd_comp   : forall ht ct c , SCND ht ct -> Rule ct [c] -> SCND ht [c]
222   | scnd_weak   : forall c       , SCND c  []
223   | scnd_leaf   : forall ht c    , SCND ht [c]  -> SCND ht [c]
224   | scnd_branch : forall ht c1 c2, SCND ht c1 -> SCND ht c2 -> SCND ht (c1,,c2)
225   .
226   Hint Constructors SCND.
227
228   (* Any ND whose primitive Rules have at most one conclusion (note that nd_prod is allowed!) can be turned into an SCND. *)
229   Definition mkSCND (all_rules_one_conclusion : forall h c1 c2, Rule h (c1,,c2) -> False)
230     : forall h x c,  ND x c -> SCND h x -> SCND h c.
231     intros h x c nd.
232     induction nd; intro k.
233       apply k.
234       apply k.
235       apply scnd_weak.
236       eapply scnd_branch; apply k.
237       inversion k; subst.
238         apply (scnd_branch _ _ _ (IHnd1 X) (IHnd2 X0)).
239       apply IHnd2.
240         apply IHnd1.
241         apply k.
242       inversion k; subst; auto.
243       inversion k; subst; auto.
244       apply scnd_branch; auto.
245       apply scnd_branch; auto.
246       inversion k; subst; inversion X; subst; auto.
247       inversion k; subst; inversion X0; subst; auto.
248       destruct c.
249         destruct o.
250         apply scnd_leaf. eapply scnd_comp. apply k. apply r.
251         apply scnd_weak.
252         set (all_rules_one_conclusion _ _ _ r) as bogus.
253           inversion bogus.
254           Defined.
255
256   (* a "ClosedND" is a proof with no open hypotheses and no multi-conclusion rules *)
257   Inductive ClosedND : Tree ??Judgment -> Type :=
258   | cnd_weak   : ClosedND []
259   | cnd_rule   : forall h c    , ClosedND h  -> Rule h c    -> ClosedND c
260   | cnd_branch : forall   c1 c2, ClosedND c1 -> ClosedND c2 -> ClosedND (c1,,c2)
261   .
262
263   (* we can turn an SCND without hypotheses into a ClosedND *)
264   Definition closedFromSCND h c (pn2:SCND h c)(cnd:ClosedND h) : ClosedND c.
265   refine ((fix closedFromPnodes h c (pn2:SCND h c)(cnd:ClosedND h) {struct pn2} := 
266     (match pn2 in SCND H C return H=h -> C=c -> _  with
267       | scnd_weak   c                 => let case_weak := tt in _
268       | scnd_leaf   ht z pn'          => let case_leaf := tt in let qq := closedFromPnodes _ _ pn' in _
269       | scnd_comp  ht ct c pn' rule   => let case_comp := tt in let qq := closedFromPnodes _ _ pn' in _
270       | scnd_branch ht c1 c2 pn' pn'' => let case_branch := tt in
271                                         let q1 := closedFromPnodes _ _ pn' in 
272                                         let q2 := closedFromPnodes _ _ pn'' in _
273
274     end (refl_equal _) (refl_equal _))) h c pn2 cnd).
275
276   destruct case_comp.
277     intros.
278     clear pn2.
279     apply (cnd_rule ct).
280     apply qq.
281     subst.
282     apply cnd0.
283     apply rule.
284
285   destruct case_weak.
286     intros; subst.
287     apply cnd_weak.
288
289   destruct case_leaf.
290     intros.
291     apply qq.
292     subst.
293     apply cnd0.
294
295   destruct case_branch.
296     intros.
297     apply cnd_branch.
298     apply q1. subst. apply cnd0.
299     apply q2. subst. apply cnd0.
300     Defined.
301
302   (* undo the above *)
303   Fixpoint closedNDtoNormalND {c}(cnd:ClosedND c) : ND [] c :=
304   match cnd in ClosedND C return ND [] C with
305   | cnd_weak                   => nd_id0
306   | cnd_rule   h c cndh rhc    => closedNDtoNormalND cndh ;; nd_rule rhc
307   | cnd_branch c1 c2 cnd1 cnd2 => nd_llecnac ;; nd_prod (closedNDtoNormalND cnd1) (closedNDtoNormalND cnd2)
308   end.
309
310   Close Scope nd_scope.
311   Open Scope pf_scope.
312
313 End Natural_Deduction.
314
315 Implicit Arguments ND [ Judgment ].
316 Hint Constructors Structural.
317 Hint Extern 1 => apply nd_id_structural.
318 Hint Extern 1 => apply ndr_structural_indistinguishable.
319
320 (* This first notation gets its own scope because it can be confusing when we're working with multiple different kinds
321  * of proofs.  When only one kind of proof is in use, it's quite helpful though. *)
322 Notation "H /⋯⋯/ C" := (@ND _ _ H C)             : pf_scope.
323 Notation "a ;; b"   := (nd_comp a b)             : nd_scope.
324 Notation "a ** b"   := (nd_prod a b)             : nd_scope.
325 Notation "[# a #]"  := (nd_rule a)               : nd_scope.
326 Notation "a === b"  := (@ndr_eqv _ _ _ _ _ a b)  : nd_scope.
327
328 (* enable setoid rewriting *)
329 Open Scope nd_scope.
330 Open Scope pf_scope.
331
332 Add Parametric Relation {jt rt ndr h c} : (h/⋯⋯/c) (@ndr_eqv jt rt ndr h c)
333   reflexivity proved by  (@Equivalence_Reflexive  _ _ (ndr_eqv_equivalence h c))
334   symmetry proved by     (@Equivalence_Symmetric  _ _ (ndr_eqv_equivalence h c))
335   transitivity proved by (@Equivalence_Transitive _ _ (ndr_eqv_equivalence h c))
336     as parametric_relation_ndr_eqv.
337   Add Parametric Morphism {jt rt ndr h x c} : (@nd_comp jt rt h x c)
338   with signature ((ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)))
339     as parametric_morphism_nd_comp.
340     intros; apply ndr_comp_respects; auto.
341     Defined.
342   Add Parametric Morphism {jt rt ndr a b c d} : (@nd_prod jt rt a b c d)
343   with signature ((ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)))
344     as parametric_morphism_nd_prod.
345     intros; apply ndr_prod_respects; auto.
346     Defined.
347
348 (* a generalization of the procedure used to build (nd_id n) from nd_id0 and nd_id1 *)
349 Definition nd_replicate
350   : forall
351     {Judgment}{Rule}{Ob}
352     (h:Ob->Judgment)
353     (c:Ob->Judgment)
354     (j:Tree ??Ob),
355     (forall (o:Ob), @ND Judgment Rule [h o] [c o]) ->
356     @ND Judgment Rule (mapOptionTree h j) (mapOptionTree c j).
357   intros.
358   induction j; simpl.
359     destruct a; simpl.
360     apply X.
361     apply nd_id0.
362     apply nd_prod; auto.
363     Defined.
364
365 (* "map" over natural deduction proofs, where the result proof has the same judgments (but different rules) *)
366 Definition nd_map
367   : forall
368     {Judgment}{Rule0}{Rule1}
369     (r:forall h c, Rule0 h c -> @ND Judgment Rule1 h c)
370     {h}{c}
371     (pf:@ND Judgment Rule0 h c)
372     ,
373     @ND Judgment Rule1 h c.
374     intros Judgment Rule0 Rule1 r.
375
376     refine ((fix nd_map h c pf {struct pf} :=
377      ((match pf
378          in @ND _ _ H C
379           return
380           @ND Judgment Rule1 H C
381         with
382         | nd_id0                     => let case_nd_id0     := tt in _
383         | nd_id1     h               => let case_nd_id1     := tt in _
384         | nd_weak    h               => let case_nd_weak    := tt in _
385         | nd_copy    h               => let case_nd_copy    := tt in _
386         | nd_prod    _ _ _ _ lpf rpf => let case_nd_prod    := tt in _
387         | nd_comp    _ _ _   top bot => let case_nd_comp    := tt in _
388         | nd_rule    _ _     rule    => let case_nd_rule    := tt in _
389         | nd_cancell _               => let case_nd_cancell := tt in _
390         | nd_cancelr _               => let case_nd_cancelr := tt in _
391         | nd_llecnac _               => let case_nd_llecnac := tt in _
392         | nd_rlecnac _               => let case_nd_rlecnac := tt in _
393         | nd_assoc   _ _ _           => let case_nd_assoc   := tt in _
394         | nd_cossa   _ _ _           => let case_nd_cossa   := tt in _
395       end))) ); simpl in *.
396
397     destruct case_nd_id0.      apply nd_id0.
398     destruct case_nd_id1.      apply nd_id1.
399     destruct case_nd_weak.     apply nd_weak.
400     destruct case_nd_copy.     apply nd_copy.
401     destruct case_nd_prod.     apply (nd_prod (nd_map _ _ lpf) (nd_map _ _ rpf)).
402     destruct case_nd_comp.     apply (nd_comp (nd_map _ _ top) (nd_map _ _ bot)).
403     destruct case_nd_cancell.  apply nd_cancell.
404     destruct case_nd_cancelr.  apply nd_cancelr.
405     destruct case_nd_llecnac.  apply nd_llecnac.
406     destruct case_nd_rlecnac.  apply nd_rlecnac.
407     destruct case_nd_assoc.    apply nd_assoc.
408     destruct case_nd_cossa.    apply nd_cossa.
409     apply r. apply rule.
410     Defined.
411
412 (* "map" over natural deduction proofs, where the result proof has different judgments *)
413 Definition nd_map'
414   : forall
415     {Judgment0}{Rule0}{Judgment1}{Rule1}
416     (f:Judgment0->Judgment1)
417     (r:forall h c, Rule0 h c -> @ND Judgment1 Rule1 (mapOptionTree f h) (mapOptionTree f c))
418     {h}{c}
419     (pf:@ND Judgment0 Rule0 h c)
420     ,
421     @ND Judgment1 Rule1 (mapOptionTree f h) (mapOptionTree f c).
422     intros Judgment0 Rule0 Judgment1 Rule1 f r.
423     
424     refine ((fix nd_map' h c pf {struct pf} :=
425      ((match pf
426          in @ND _ _ H C
427           return
428           @ND Judgment1 Rule1 (mapOptionTree f H) (mapOptionTree f C)
429         with
430         | nd_id0                     => let case_nd_id0     := tt in _
431         | nd_id1     h               => let case_nd_id1     := tt in _
432         | nd_weak    h               => let case_nd_weak    := tt in _
433         | nd_copy    h               => let case_nd_copy    := tt in _
434         | nd_prod    _ _ _ _ lpf rpf => let case_nd_prod    := tt in _
435         | nd_comp    _ _ _   top bot => let case_nd_comp    := tt in _
436         | nd_rule    _ _     rule    => let case_nd_rule    := tt in _
437         | nd_cancell _               => let case_nd_cancell := tt in _
438         | nd_cancelr _               => let case_nd_cancelr := tt in _
439         | nd_llecnac _               => let case_nd_llecnac := tt in _
440         | nd_rlecnac _               => let case_nd_rlecnac := tt in _
441         | nd_assoc   _ _ _           => let case_nd_assoc   := tt in _
442         | nd_cossa   _ _ _           => let case_nd_cossa   := tt in _
443       end))) ); simpl in *.
444
445     destruct case_nd_id0.      apply nd_id0.
446     destruct case_nd_id1.      apply nd_id1.
447     destruct case_nd_weak.     apply nd_weak.
448     destruct case_nd_copy.     apply nd_copy.
449     destruct case_nd_prod.     apply (nd_prod (nd_map' _ _ lpf) (nd_map' _ _ rpf)).
450     destruct case_nd_comp.     apply (nd_comp (nd_map' _ _ top) (nd_map' _ _ bot)).
451     destruct case_nd_cancell.  apply nd_cancell.
452     destruct case_nd_cancelr.  apply nd_cancelr.
453     destruct case_nd_llecnac.  apply nd_llecnac.
454     destruct case_nd_rlecnac.  apply nd_rlecnac.
455     destruct case_nd_assoc.    apply nd_assoc.
456     destruct case_nd_cossa.    apply nd_cossa.
457     apply r. apply rule.
458     Defined.
459
460 (* witnesses the fact that every Rule in a particular proof satisfies the given predicate *)
461 Inductive nd_property {Judgment}{Rule}(P:forall h c, @Rule h c -> Prop) : forall {h}{c}, @ND Judgment Rule h c -> Prop :=
462   | nd_property_structural      : forall h c pf, Structural pf -> @nd_property _ _ P h c pf
463   | nd_property_prod            : forall h0 c0 pf0 h1 c1 pf1,
464     @nd_property _ _ P h0 c0 pf0 -> @nd_property _ _ P h1 c1 pf1 -> @nd_property _ _ P _ _ (nd_prod pf0 pf1)
465   | nd_property_comp            : forall h0 c0 pf0  c1 pf1,
466     @nd_property _ _ P h0 c0 pf0 -> @nd_property _ _ P c0 c1 pf1 -> @nd_property _ _ P _ _ (nd_comp pf0 pf1)
467   | nd_property_rule            : forall h c r, P h c r -> @nd_property _ _ P h c (nd_rule r).
468   Hint Constructors nd_property.
469
470 Close Scope pf_scope.
471 Close Scope nd_scope.