56d74cda5ecb82dba4d32a8f12e3bf24d701edab
[coq-hetmet.git] / src / NaturalDeduction.v
1 (*********************************************************************************************************************************)
2 (* NaturalDeduction:                                                                                                             *)
3 (*                                                                                                                               *)
4 (*   Structurally explicit natural deduction proofs.                                                                             *)
5 (*                                                                                                                               *)
6 (*********************************************************************************************************************************)
7
8 Generalizable All Variables.
9 Require Import Preamble.
10 Require Import General.
11 Require Import Coq.Strings.Ascii.
12 Require Import Coq.Strings.String.
13
14 (*
15  * Unlike most formalizations, this library offers two different ways
16  * to represent a natural deduction proof.  To demonstrate this,
17  * consider the signature of the propositional calculus:
18  *
19  *   Variable  PropositionalVariable : Type.
20  *
21  *   Inductive Formula : Prop :=
22  *   | formula_var : PropositionalVariable -> Formula   (* every propositional variable is a formula *)
23  *   | formula_and :   Formula ->  Formula -> Formula   (* the conjunction of any two formulae is a formula *)
24  *   | formula_or  :   Formula ->  Formula -> Formula   (* the disjunction of any two formulae is a formula *)
25  *
26  * And couple this with the theory of conjunction and disjunction:
27  * φ\/ψ is true if either φ is true or ψ is true, and φ/\ψ is true
28  * if both φ and ψ are true.
29  *
30  * 1) Structurally implicit proofs
31  *
32  *    This is what you would call the "usual" representation –- it's
33  *    what most people learn when they first start programming in Coq:
34  *
35  *    Inductive IsTrue : Formula -> Prop :=
36  *    | IsTrue_or1 : forall f1 f2, IsTrue f1 ->              IsTrue (formula_or  f1 f2) 
37  *    | IsTrue_or2 : forall f1 f2,              IsTrue f2 -> IsTrue (formula_or  f1 f2) 
38  *    | IsTrue_and : forall f1 f2, IsTrue f2 -> IsTrue f2 -> IsTrue (formula_and f1 f2) 
39  *
40  *    Here each judgment (such as "φ is true") is represented by a Coq
41  *    type; furthermore:
42  *
43  *       1. A proof of a judgment is any inhabitant of that Coq type.
44  *
45  *       2. A proof of a judgment "J2" from hypothesis judgment "J1"
46  *          is any Coq function from the Coq type for J1 to the Coq
47  *          type for J2; Composition of (hypothetical) proofs is
48  *          represented by composition of Coq functions.
49  *
50  *       3. A pair of judgments is represented by their product (Coq
51  *          type [prod]) in Coq; a pair of proofs is represented by
52  *          their pair (Coq function [pair]) in Coq.
53  *
54  *       4. Duplication of hypotheses is represented by the Coq
55  *          function (fun x => (x,x)).  Dereliction of hypotheses is
56  *          represented by the coq function (fun (x,y) => x) or (fun
57  *          (x,y) => y).  Exchange of the order of hypotheses is
58  *          represented by the Coq function (fun (x,y) => (y,x)).
59  *
60  *    The above can be done using lists instead of tuples.
61  *
62  *    The advantage of this approach is that it requires a minimum
63  *    amount of syntax, and takes maximum advantage of Coq's
64  *    automation facilities.
65  *
66  *    The disadvantage is that one cannot reason about proof-theoretic
67  *    properties *generically* across different signatures and
68  *    theories.  Each signature has its own type of judgments, and
69  *    each theory has its own type of proofs.  In the present
70  *    development we will want to prove –– in this generic manner --
71  *    that various classes of natural deduction calculi form various
72  *    kinds of categories.  So we will need this ability to reason
73  *    about proofs independently of the type used to represent
74  *    judgments and (more importantly) of the set of basic inference
75  *    rules.
76  *
77  * 2) Structurally explicit proofs
78  *
79  *    Structurally explicit proofs are formalized in this file
80  *    (NaturalDeduction.v) and are designed specifically in order to
81  *    circumvent the problem in the previous paragraph.
82  *
83  *    These proofs are actually structurally explicit on (potentially)
84  *    two different levels.  The beginning of this file formalizes
85  *    natural deduction proofs with explicit structural operations for
86  *    manipulating lists of judgments – for example, the open
87  *    hypotheses of an incomplete proof.  The class
88  *    TreeStructuralRules further down in the file instantiates ND
89  *    such that Judgments is actually a pair of trees of propositions,
90  *    and there will be a whole *other* set of rules for manipulating
91  *    the structure of a tree of propositions *within* a single
92  *    judgment.
93  *
94  *    The flattening functor ends up mapping the first kind of
95  *    structural operation (moving around judgments) onto the second
96  *    kind (moving around propositions/types).  That's why everything
97  *    is so laboriously explicit - there's important information in
98  *    those structural operations.
99  *)
100
101 (*
102  * REGARDING LISTS versus TREES:
103  *
104  * You'll notice that this formalization uses (Tree (option A)) in a
105  * lot of places where you might find (list A) more natural.  If this
106  * bothers you, see the end of the file for the technical reasons why.
107  * In short, it lets us avoid having to mess about with JMEq or EqDep,
108  * which are not as well-supported by the Coq core as the theory of
109  * CiC proper.
110  *)
111
112 Section Natural_Deduction.
113
114   (* any Coq Type may be used as the set of judgments *)
115   Context {Judgment : Type}.
116
117   (* any Coq Type –- indexed by the hypothesis and conclusion judgments -- may be used as the set of basic inference rules *)
118   Context {Rule     : forall (hypotheses:Tree ??Judgment)(conclusion:Tree ??Judgment), Type}.
119
120   (*
121    *  This type represents a valid Natural Deduction proof from a list
122    *  (tree) of hypotheses; the notation H/⋯⋯/C is meant to look like
123    *  a proof tree with the middle missing if you tilt your head to
124    *  the left (yeah, I know it's a stretch).  Note also that this
125    *  type is capable of representing proofs with multiple
126    *  conclusions, whereas a Rule may have only one conclusion.
127    *) 
128   Inductive ND :
129     forall hypotheses:Tree ??Judgment,
130       forall conclusions:Tree ??Judgment,
131         Type :=
132
133     (* natural deduction: you may infer nothing from nothing *)
134     | nd_id0    :             [   ] /⋯⋯/ [   ]
135
136     (* natural deduction: you may infer anything from itself -- "identity proof" *)
137     | nd_id1    : forall  h,  [ h ] /⋯⋯/ [ h ]
138   
139     (* natural deduction: you may discard conclusions *)
140     | nd_weak1  : forall  h,  [ h ] /⋯⋯/ [   ]
141   
142     (* natural deduction: you may duplicate conclusions *)
143     | nd_copy   : forall  h,    h   /⋯⋯/ (h,,h)
144   
145     (* natural deduction: you may write two proof trees side by side on a piece of paper -- "proof product" *)
146     | nd_prod : forall {h1 h2 c1 c2}
147        (pf1: h1       /⋯⋯/ c1      )
148        (pf2:       h2 /⋯⋯/       c2),
149        (     h1 ,, h2 /⋯⋯/ c1 ,, c2)
150   
151     (* natural deduction: given a proof of every hypothesis, you may discharge them -- "proof composition" *)
152     | nd_comp :
153       forall {h x c}
154       `(pf1: h /⋯⋯/ x)
155       `(pf2: x /⋯⋯/ c),
156        (     h /⋯⋯/ c)
157   
158     (* Structural rules on lists of judgments - note that this is completely separate from the structural
159      * rules for *contexts* within a sequent.  The rules below manipulate lists of *judgments* rather than
160      * lists of *propositions*. *)
161     | nd_cancell : forall {a},       [] ,, a /⋯⋯/ a
162     | nd_cancelr : forall {a},       a ,, [] /⋯⋯/ a
163     | nd_llecnac : forall {a},             a /⋯⋯/ [] ,, a
164     | nd_rlecnac : forall {a},             a /⋯⋯/ a ,, []
165     | nd_assoc   : forall {a b c}, (a,,b),,c /⋯⋯/ a,,(b,,c)
166     | nd_cossa   : forall {a b c}, a,,(b,,c) /⋯⋯/ (a,,b),,c
167
168     (* any Rule by itself counts as a proof *)
169     | nd_rule    : forall {h c} (r:Rule h c), h /⋯⋯/ c
170   
171     where  "H /⋯⋯/ C" := (ND H C).
172
173     Notation "H /⋯⋯/ C" := (ND H C) : pf_scope.
174     Notation "a ;; b"   := (nd_comp a b) : nd_scope.
175     Notation "a ** b"   := (nd_prod a b) : nd_scope.
176     Open Scope nd_scope.
177     Open Scope pf_scope.
178
179   (* a predicate on proofs *)
180   Definition NDPredicate := forall h c, h /⋯⋯/ c -> Prop.
181
182   (* the structural inference rules are those which do not change, add, remove, or re-order the judgments *)
183   Inductive Structural : forall {h c}, h /⋯⋯/ c -> Prop :=
184   | nd_structural_id0     :                                                                            Structural nd_id0
185   | nd_structural_id1     : forall h,                                                                  Structural (nd_id1 h)
186   | nd_structural_cancell : forall {a},                                                                Structural (@nd_cancell a)
187   | nd_structural_cancelr : forall {a},                                                                Structural (@nd_cancelr a)
188   | nd_structural_llecnac : forall {a},                                                                Structural (@nd_llecnac a)
189   | nd_structural_rlecnac : forall {a},                                                                Structural (@nd_rlecnac a)
190   | nd_structural_assoc   : forall {a b c},                                                            Structural (@nd_assoc a b c)
191   | nd_structural_cossa   : forall {a b c},                                                            Structural (@nd_cossa a b c)
192   .
193
194   (* the closure of an NDPredicate under nd_comp and nd_prod *)
195   Inductive NDPredicateClosure (P:NDPredicate) : forall {h c}, h /⋯⋯/ c -> Prop :=
196   | ndpc_p       : forall h c f, P h c f                                                   -> NDPredicateClosure P f
197   | ndpc_prod    : forall `(pf1:h1/⋯⋯/c1)`(pf2:h2/⋯⋯/c2),
198     NDPredicateClosure P pf1 -> NDPredicateClosure P pf2 -> NDPredicateClosure P (pf1**pf2)
199   | ndpc_comp    : forall `(pf1:h1/⋯⋯/x) `(pf2: x/⋯⋯/c2),
200     NDPredicateClosure P pf1 -> NDPredicateClosure P pf2 -> NDPredicateClosure P (pf1;;pf2).
201
202   (* proofs built up from structural rules via comp and prod *)
203   Definition StructuralND {h}{c} f := @NDPredicateClosure (@Structural) h c f.
204
205   (* The Predicate (BuiltFrom f P h) asserts that "h" was built from a single occurrence of "f" and proofs which satisfy P *)
206   Inductive BuiltFrom {h'}{c'}(f:h'/⋯⋯/c')(P:NDPredicate) : forall {h c}, h/⋯⋯/c -> Prop :=
207   | builtfrom_refl  : BuiltFrom f P f
208   | builtfrom_P     : forall h c g, @P h c g -> BuiltFrom f P g
209   | builtfrom_prod1 : forall h1 c1 f1 h2 c2 f2, P h1 c1 f1 -> @BuiltFrom _ _ f P h2 c2 f2 -> BuiltFrom f P (f1 ** f2)
210   | builtfrom_prod2 : forall h1 c1 f1 h2 c2 f2, P h1 c1 f1 -> @BuiltFrom _ _ f P h2 c2 f2 -> BuiltFrom f P (f2 ** f1)
211   | builtfrom_comp1 : forall h x c       f1 f2, P h  x  f1 -> @BuiltFrom _ _ f P x  c  f2 -> BuiltFrom f P (f1 ;; f2)
212   | builtfrom_comp2 : forall h x c       f1 f2, P x  c  f1 -> @BuiltFrom _ _ f P h  x  f2 -> BuiltFrom f P (f2 ;; f1).
213
214   (* multi-judgment generalization of nd_id0 and nd_id1; making nd_id0/nd_id1 primitive and deriving all other *)
215   Fixpoint nd_id (sl:Tree ??Judgment) : sl /⋯⋯/ sl :=
216     match sl with
217       | T_Leaf None      => nd_id0
218       | T_Leaf (Some x)  => nd_id1 x
219       | T_Branch a b     => nd_prod (nd_id a) (nd_id b)
220     end.
221
222   Fixpoint nd_weak (sl:Tree ??Judgment) : sl /⋯⋯/ [] :=
223     match sl as SL return SL /⋯⋯/ [] with
224       | T_Leaf None      => nd_id0
225       | T_Leaf (Some x)  => nd_weak1 x
226       | T_Branch a b     => nd_prod (nd_weak a) (nd_weak b) ;; nd_cancelr
227     end.
228
229   Hint Constructors Structural.
230   Hint Constructors BuiltFrom.
231   Hint Constructors NDPredicateClosure.
232
233   Hint Extern 1 => apply nd_structural_id0.     
234   Hint Extern 1 => apply nd_structural_id1.     
235   Hint Extern 1 => apply nd_structural_cancell. 
236   Hint Extern 1 => apply nd_structural_cancelr. 
237   Hint Extern 1 => apply nd_structural_llecnac. 
238   Hint Extern 1 => apply nd_structural_rlecnac. 
239   Hint Extern 1 => apply nd_structural_assoc.   
240   Hint Extern 1 => apply nd_structural_cossa.   
241   Hint Extern 1 => apply ndpc_p.
242   Hint Extern 1 => apply ndpc_prod.
243   Hint Extern 1 => apply ndpc_comp.
244
245   Lemma nd_id_structural : forall sl, StructuralND (nd_id sl).
246     intros.
247     induction sl; simpl; auto.
248     destruct a; auto.
249     Defined.
250
251   (* An equivalence relation on proofs which is sensitive only to the logical content of the proof -- insensitive to
252    * structural variations  *)
253   Class ND_Relation :=
254   { ndr_eqv                  : forall {h c  }, h /⋯⋯/ c -> h /⋯⋯/ c -> Prop where "pf1 === pf2" := (@ndr_eqv _ _  pf1 pf2)
255   ; ndr_eqv_equivalence      : forall h c, Equivalence (@ndr_eqv h c)
256
257   (* the relation must respect composition, be associative wrt composition, and be left and right neutral wrt the identity proof *)
258   ; ndr_comp_respects        : forall {a b c}(f f':a/⋯⋯/b)(g g':b/⋯⋯/c),      f === f' -> g === g' -> f;;g === f';;g'
259   ; ndr_comp_associativity   : forall `(f:a/⋯⋯/b)`(g:b/⋯⋯/c)`(h:c/⋯⋯/d),                         (f;;g);;h === f;;(g;;h)
260
261   (* the relation must respect products, be associative wrt products, and be left and right neutral wrt the identity proof *)
262   ; ndr_prod_respects        : forall {a b c d}(f f':a/⋯⋯/b)(g g':c/⋯⋯/d),     f===f' -> g===g' ->    f**g === f'**g'
263   ; ndr_prod_associativity   : forall `(f:a/⋯⋯/a')`(g:b/⋯⋯/b')`(h:c/⋯⋯/c'),       (f**g)**h === nd_assoc ;; f**(g**h) ;; nd_cossa
264
265   (* products and composition must distribute over each other *)
266   ; 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')
267
268   (* Given a proof f, any two proofs built from it using only structural rules are indistinguishable.  Keep in mind that
269    * nd_weak and nd_copy aren't considered structural, so the hypotheses and conclusions of such proofs will be an identical
270    * list, differing only in the "parenthesization" and addition or removal of empty leaves. *)
271   ; ndr_builtfrom_structural : forall `(f:a/⋯⋯/b){a' b'}(g1 g2:a'/⋯⋯/b'),
272     BuiltFrom f (@StructuralND) g1 ->
273     BuiltFrom f (@StructuralND) g2 ->
274     g1 === g2
275
276   (* proofs of nothing are not distinguished from each other *)
277   ; ndr_void_proofs_irrelevant : forall `(f:a/⋯⋯/[])(g:a/⋯⋯/[]), f === g
278
279   (* products and duplication must distribute over each other *)
280   ; ndr_prod_preserves_copy  : forall `(f:a/⋯⋯/b),                                        nd_copy a;; f**f === f ;; nd_copy b
281
282   (* duplicating a hypothesis and discarding it is irrelevant *)
283   ; ndr_copy_then_weak_left   : forall a,                            nd_copy a;; (nd_weak _ ** nd_id _) ;; nd_cancell === nd_id _
284   ; ndr_copy_then_weak_right  : forall a,                            nd_copy a;; (nd_id _ ** nd_weak _) ;; nd_cancelr === nd_id _
285   }.
286
287   (* 
288    * Natural Deduction proofs which are Structurally Implicit on the
289    * level of judgments.  These proofs have poor compositionality
290    * properties (vertically, they look more like lists than trees) but
291    * are easier to do induction over.
292    *)
293   Inductive SIND : Tree ??Judgment -> Tree ??Judgment -> Type :=
294   | scnd_weak   : forall c       , SIND c  []
295   | scnd_comp   : forall ht ct c , SIND ht ct -> Rule ct [c] -> SIND ht [c]
296   | scnd_branch : forall ht c1 c2, SIND ht c1 -> SIND ht c2 -> SIND ht (c1,,c2)
297   .
298   Hint Constructors SIND.
299
300   (* Any ND whose primitive Rules have at most one conclusion (note that nd_prod is allowed!) can be turned into an SIND. *)
301   Definition mkSIND (all_rules_one_conclusion : forall h c1 c2, Rule h (c1,,c2) -> False)
302     : forall h x c,  ND x c -> SIND h x -> SIND h c.
303     intros h x c nd.
304     induction nd; intro k.
305       apply k.
306       apply k.
307       apply scnd_weak.
308       eapply scnd_branch; apply k.
309       inversion k; subst.
310         apply (scnd_branch _ _ _ (IHnd1 X) (IHnd2 X0)).
311       apply IHnd2.
312         apply IHnd1.
313         apply k.
314       inversion k; subst; auto.
315       inversion k; subst; auto.
316       apply scnd_branch; auto.
317       apply scnd_branch; auto.
318       inversion k; subst; inversion X; subst; auto.
319       inversion k; subst; inversion X0; subst; auto.
320       destruct c.
321         destruct o.
322         eapply scnd_comp. apply k. apply r.
323         apply scnd_weak.
324         set (all_rules_one_conclusion _ _ _ r) as bogus.
325           inversion bogus.
326           Defined.
327
328   (* a "ClosedSIND" is a proof with no open hypotheses and no multi-conclusion rules *)
329   Inductive ClosedSIND : Tree ??Judgment -> Type :=
330   | cnd_weak   : ClosedSIND []
331   | cnd_rule   : forall h c    , ClosedSIND h  -> Rule h c    -> ClosedSIND c
332   | cnd_branch : forall   c1 c2, ClosedSIND c1 -> ClosedSIND c2 -> ClosedSIND (c1,,c2)
333   .
334
335   (* we can turn an SIND without hypotheses into a ClosedSIND *)
336   Definition closedFromSIND h c (pn2:SIND h c)(cnd:ClosedSIND h) : ClosedSIND c.
337   refine ((fix closedFromPnodes h c (pn2:SIND h c)(cnd:ClosedSIND h) {struct pn2} := 
338     (match pn2 in SIND H C return H=h -> C=c -> _  with
339       | scnd_weak   c                 => let case_weak := tt in _
340       | scnd_comp  ht ct c pn' rule   => let case_comp := tt in let qq := closedFromPnodes _ _ pn' in _
341       | scnd_branch ht c1 c2 pn' pn'' => let case_branch := tt in
342                                          let q1 := closedFromPnodes _ _ pn' in 
343                                          let q2 := closedFromPnodes _ _ pn'' in _
344
345     end (refl_equal _) (refl_equal _))) h c pn2 cnd).
346
347   destruct case_weak.
348     intros; subst.
349     apply cnd_weak.
350
351   destruct case_comp.
352     intros.
353     clear pn2.
354     apply (cnd_rule ct).
355     apply qq.
356     subst.
357     apply cnd0.
358     apply rule.
359
360   destruct case_branch.
361     intros.
362     apply cnd_branch.
363     apply q1. subst. apply cnd0.
364     apply q2. subst. apply cnd0.
365     Defined.
366
367   (* undo the above *)
368   Fixpoint closedNDtoNormalND {c}(cnd:ClosedSIND c) : ND [] c :=
369   match cnd in ClosedSIND C return ND [] C with
370   | cnd_weak                   => nd_id0
371   | cnd_rule   h c cndh rhc    => closedNDtoNormalND cndh ;; nd_rule rhc
372   | cnd_branch c1 c2 cnd1 cnd2 => nd_llecnac ;; nd_prod (closedNDtoNormalND cnd1) (closedNDtoNormalND cnd2)
373   end.
374
375   (* Natural Deduction systems whose judgments happen to be pairs of the same type *)
376   Section SequentND.
377     Context {S:Type}.                   (* type of sequent components *)
378     Context {sequent:S->S->Judgment}.   (* pairing operation which forms a sequent from its halves *)
379     Notation "a |= b" := (sequent a b).
380
381     (* a SequentND is a natural deduction whose judgments are sequents, has initial sequents, and has a cut rule *)
382     Class SequentND :=
383     { snd_initial : forall a,                           [ ] /⋯⋯/ [ a |= a ]
384     ; snd_cut     : forall a b c,  [ a |= b ] ,, [ b |= c ] /⋯⋯/ [ a |= c ]
385     }.
386
387     Context (sequentND:SequentND).
388     Context (ndr:ND_Relation).
389
390     (*
391      * A predicate singling out structural rules, initial sequents,
392      * and cut rules.
393      *
394      * Proofs using only structural rules cannot add or remove
395      * judgments - their hypothesis and conclusion judgment-trees will
396      * differ only in "parenthesization" and the presence/absence of
397      * empty leaves.  This means that a proof involving only
398      * structural rules, cut, and initial sequents can ADD new
399      * non-empty judgment-leaves only via snd_initial, and can only
400      * REMOVE non-empty judgment-leaves only via snd_cut.  Since the
401      * initial sequent is a left and right identity for cut, and cut
402      * is associative, any two proofs (with the same hypotheses and
403      * conclusions) using only structural rules, cut, and initial
404      * sequents are logically indistinguishable - their differences
405      * are logically insignificant.
406      *
407      * Note that it is important that nd_weak and nd_copy aren't
408      * considered to be "structural".
409      *)
410     Inductive SequentND_Inert : forall h c, h/⋯⋯/c -> Prop :=
411     | snd_inert_initial   : forall a,                     SequentND_Inert _ _ (snd_initial a)
412     | snd_inert_cut       : forall a b c,                 SequentND_Inert _ _ (snd_cut a b c)
413     | snd_inert_structural: forall a b f, Structural f -> SequentND_Inert a b f
414     .
415
416     (* An ND_Relation for a sequent deduction should not distinguish between two proofs having the same hypotheses and conclusions
417      * if those proofs use only initial sequents, cut, and structural rules (see comment above) *)
418     Class SequentND_Relation :=
419     { sndr_ndr   := ndr
420     ; sndr_inert : forall a b (f g:a/⋯⋯/b),
421       NDPredicateClosure SequentND_Inert f -> 
422       NDPredicateClosure SequentND_Inert g -> 
423       ndr_eqv f g }.
424
425   End SequentND.
426
427   (* Deductions on sequents whose antecedent is a tree of propositions (i.e. a context) *)
428   Section ContextND.
429     Context {P:Type}{sequent:Tree ??P -> Tree ??P -> Judgment}.
430     Context {snd:SequentND(sequent:=sequent)}.
431     Notation "a |= b" := (sequent a b).
432
433     (* Note that these rules mirror nd_{cancell,cancelr,rlecnac,llecnac,assoc,cossa} but are completely separate from them *)
434     Class ContextND :=
435     { cnd_ant_assoc     : forall x a b c, ND [((a,,b),,c) |= x]     [(a,,(b,,c)) |= x   ]
436     ; cnd_ant_cossa     : forall x a b c, ND [(a,,(b,,c)) |= x]     [((a,,b),,c) |= x   ]
437     ; cnd_ant_cancell   : forall x a    , ND [  [],,a     |= x]     [        a   |= x   ]
438     ; cnd_ant_cancelr   : forall x a    , ND [a,,[]       |= x]     [        a   |= x   ]
439     ; cnd_ant_llecnac   : forall x a    , ND [      a     |= x]     [    [],,a   |= x   ]
440     ; cnd_ant_rlecnac   : forall x a    , ND [      a     |= x]     [    a,,[]   |= x   ]
441     ; cnd_expand_left   : forall a b c  , ND [          a |= b]     [ c,,a       |= c,,b]
442     ; cnd_expand_right  : forall a b c  , ND [          a |= b]     [ a,,c       |= b,,c]
443     ; cnd_snd           := snd
444     }.
445
446     Context `(ContextND).
447
448     (*
449      * A predicate singling out initial sequents, cuts, context expansion,
450      * and structural rules.
451      *
452      * Any two proofs (with the same hypotheses and conclusions) whose
453      * non-structural rules do nothing other than expand contexts,
454      * re-arrange contexts, or introduce additional initial-sequent
455      * conclusions are indistinguishable.  One important consequence
456      * is that asking for a small initial sequent and then expanding
457      * it using cnd_expand_{right,left} is no different from simply
458      * asking for the larger initial sequent in the first place.
459      *
460      *)
461     Inductive ContextND_Inert : forall h c, h/⋯⋯/c -> Prop :=
462     | cnd_inert_initial           : forall a,                     ContextND_Inert _ _ (snd_initial a)
463     | cnd_inert_cut               : forall a b c,                 ContextND_Inert _ _ (snd_cut a b c)
464     | cnd_inert_structural        : forall a b f, Structural f -> ContextND_Inert a b f
465     | cnd_inert_cnd_ant_assoc     : forall x a b c, ContextND_Inert _ _ (cnd_ant_assoc   x a b c)
466     | cnd_inert_cnd_ant_cossa     : forall x a b c, ContextND_Inert _ _ (cnd_ant_cossa   x a b c)
467     | cnd_inert_cnd_ant_cancell   : forall x a    , ContextND_Inert _ _ (cnd_ant_cancell x a)
468     | cnd_inert_cnd_ant_cancelr   : forall x a    , ContextND_Inert _ _ (cnd_ant_cancelr x a)
469     | cnd_inert_cnd_ant_llecnac   : forall x a    , ContextND_Inert _ _ (cnd_ant_llecnac x a)
470     | cnd_inert_cnd_ant_rlecnac   : forall x a    , ContextND_Inert _ _ (cnd_ant_rlecnac x a)
471     | cnd_inert_se_expand_left    : forall t g s  , ContextND_Inert _ _ (@cnd_expand_left  _ t g s)
472     | cnd_inert_se_expand_right   : forall t g s  , ContextND_Inert _ _ (@cnd_expand_right _ t g s).
473
474     Class ContextND_Relation {ndr}{sndr:SequentND_Relation _ ndr} :=
475     { cndr_inert : forall {a}{b}(f g:a/⋯⋯/b),
476                            NDPredicateClosure ContextND_Inert f -> 
477                            NDPredicateClosure ContextND_Inert g -> 
478                            ndr_eqv f g
479     ; cndr_sndr  := sndr
480     }.
481
482     (* a proof is Analytic if it does not use cut *)
483     (*
484     Definition Analytic_Rule : NDPredicate :=
485       fun h c f => forall c, not (snd_cut _ _ c = f).
486     Definition AnalyticND := NDPredicateClosure Analytic_Rule.
487
488     (* a proof system has cut elimination if, for every proof, there is an analytic proof with the same conclusion *)
489     Class CutElimination :=
490     { ce_eliminate : forall h c, h/⋯⋯/c -> h/⋯⋯/c
491     ; ce_analytic  : forall h c f, AnalyticND (ce_eliminate h c f)
492     }.
493
494     (* cut elimination is strong if the analytic proof is furthermore equivalent to the original proof *)
495     Class StrongCutElimination :=
496     { sce_ce       <: CutElimination
497     ; ce_strong     : forall h c f, f === ce_eliminate h c f
498     }.
499     *)
500
501   End ContextND.
502
503   Close Scope nd_scope.
504   Open Scope pf_scope.
505
506 End Natural_Deduction.
507
508 Coercion snd_cut   : SequentND >-> Funclass.
509 Coercion cnd_snd   : ContextND >-> SequentND.
510 Coercion sndr_ndr  : SequentND_Relation >-> ND_Relation.
511 Coercion cndr_sndr : ContextND_Relation >-> SequentND_Relation.
512
513 Implicit Arguments ND [ Judgment ].
514 Hint Constructors Structural.
515 Hint Extern 1 => apply nd_id_structural.
516 Hint Extern 1 => apply ndr_builtfrom_structural.
517 Hint Extern 1 => apply nd_structural_id0.     
518 Hint Extern 1 => apply nd_structural_id1.     
519 Hint Extern 1 => apply nd_structural_cancell. 
520 Hint Extern 1 => apply nd_structural_cancelr. 
521 Hint Extern 1 => apply nd_structural_llecnac. 
522 Hint Extern 1 => apply nd_structural_rlecnac. 
523 Hint Extern 1 => apply nd_structural_assoc.   
524 Hint Extern 1 => apply nd_structural_cossa.   
525 Hint Extern 1 => apply ndpc_p.
526 Hint Extern 1 => apply ndpc_prod.
527 Hint Extern 1 => apply ndpc_comp.
528 Hint Extern 1 => apply builtfrom_refl.
529 Hint Extern 1 => apply builtfrom_prod1.
530 Hint Extern 1 => apply builtfrom_prod2.
531 Hint Extern 1 => apply builtfrom_comp1.
532 Hint Extern 1 => apply builtfrom_comp2.
533 Hint Extern 1 => apply builtfrom_P.
534
535 Hint Extern 1 => apply snd_inert_initial.
536 Hint Extern 1 => apply snd_inert_cut.
537 Hint Extern 1 => apply snd_inert_structural.
538
539 Hint Extern 1 => apply cnd_inert_initial.
540 Hint Extern 1 => apply cnd_inert_cut.
541 Hint Extern 1 => apply cnd_inert_structural.
542 Hint Extern 1 => apply cnd_inert_cnd_ant_assoc.
543 Hint Extern 1 => apply cnd_inert_cnd_ant_cossa.
544 Hint Extern 1 => apply cnd_inert_cnd_ant_cancell.
545 Hint Extern 1 => apply cnd_inert_cnd_ant_cancelr.
546 Hint Extern 1 => apply cnd_inert_cnd_ant_llecnac.
547 Hint Extern 1 => apply cnd_inert_cnd_ant_rlecnac.
548 Hint Extern 1 => apply cnd_inert_se_expand_left.
549 Hint Extern 1 => apply cnd_inert_se_expand_right.
550
551 (* This first notation gets its own scope because it can be confusing when we're working with multiple different kinds
552  * of proofs.  When only one kind of proof is in use, it's quite helpful though. *)
553 Notation "H /⋯⋯/ C" := (@ND _ _ H C)             : pf_scope.
554 Notation "a ;; b"   := (nd_comp a b)             : nd_scope.
555 Notation "a ** b"   := (nd_prod a b)             : nd_scope.
556 Notation "[# a #]"  := (nd_rule a)               : nd_scope.
557 Notation "a === b"  := (@ndr_eqv _ _ _ _ _ a b)  : nd_scope.
558
559 (* enable setoid rewriting *)
560 Open Scope nd_scope.
561 Open Scope pf_scope.
562
563 Add Parametric Relation {jt rt ndr h c} : (h/⋯⋯/c) (@ndr_eqv jt rt ndr h c)
564   reflexivity proved by  (@Equivalence_Reflexive  _ _ (ndr_eqv_equivalence h c))
565   symmetry proved by     (@Equivalence_Symmetric  _ _ (ndr_eqv_equivalence h c))
566   transitivity proved by (@Equivalence_Transitive _ _ (ndr_eqv_equivalence h c))
567     as parametric_relation_ndr_eqv.
568   Add Parametric Morphism {jt rt ndr h x c} : (@nd_comp jt rt h x c)
569   with signature ((ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)))
570     as parametric_morphism_nd_comp.
571     intros; apply ndr_comp_respects; auto.
572     Defined.
573   Add Parametric Morphism {jt rt ndr a b c d} : (@nd_prod jt rt a b c d)
574   with signature ((ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)))
575     as parametric_morphism_nd_prod.
576     intros; apply ndr_prod_respects; auto.
577     Defined.
578
579 Section ND_Relation_Facts.
580   Context `{ND_Relation}.
581
582   (* useful *)
583   Lemma ndr_comp_right_identity : forall h c (f:h/⋯⋯/c), ndr_eqv (f ;; nd_id c) f.
584     intros; apply (ndr_builtfrom_structural f); auto.
585     Defined.
586
587   (* useful *)
588   Lemma ndr_comp_left_identity : forall h c (f:h/⋯⋯/c), ndr_eqv (nd_id h ;; f) f.
589     intros; apply (ndr_builtfrom_structural f); auto.
590     Defined.
591
592 End ND_Relation_Facts.
593
594 (* a generalization of the procedure used to build (nd_id n) from nd_id0 and nd_id1 *)
595 Definition nd_replicate
596   : forall
597     {Judgment}{Rule}{Ob}
598     (h:Ob->Judgment)
599     (c:Ob->Judgment)
600     (j:Tree ??Ob),
601     (forall (o:Ob), @ND Judgment Rule [h o] [c o]) ->
602     @ND Judgment Rule (mapOptionTree h j) (mapOptionTree c j).
603   intros.
604   induction j; simpl.
605     destruct a; simpl.
606     apply X.
607     apply nd_id0.
608     apply nd_prod; auto.
609     Defined.
610
611 (* "map" over natural deduction proofs, where the result proof has the same judgments (but different rules) *)
612 Definition nd_map
613   : forall
614     {Judgment}{Rule0}{Rule1}
615     (r:forall h c, Rule0 h c -> @ND Judgment Rule1 h c)
616     {h}{c}
617     (pf:@ND Judgment Rule0 h c)
618     ,
619     @ND Judgment Rule1 h c.
620     intros Judgment Rule0 Rule1 r.
621
622     refine ((fix nd_map h c pf {struct pf} :=
623      ((match pf
624          in @ND _ _ H C
625           return
626           @ND Judgment Rule1 H C
627         with
628         | nd_id0                     => let case_nd_id0     := tt in _
629         | nd_id1     h               => let case_nd_id1     := tt in _
630         | nd_weak1   h               => let case_nd_weak    := tt in _
631         | nd_copy    h               => let case_nd_copy    := tt in _
632         | nd_prod    _ _ _ _ lpf rpf => let case_nd_prod    := tt in _
633         | nd_comp    _ _ _   top bot => let case_nd_comp    := tt in _
634         | nd_rule    _ _     rule    => let case_nd_rule    := tt in _
635         | nd_cancell _               => let case_nd_cancell := tt in _
636         | nd_cancelr _               => let case_nd_cancelr := tt in _
637         | nd_llecnac _               => let case_nd_llecnac := tt in _
638         | nd_rlecnac _               => let case_nd_rlecnac := tt in _
639         | nd_assoc   _ _ _           => let case_nd_assoc   := tt in _
640         | nd_cossa   _ _ _           => let case_nd_cossa   := tt in _
641       end))) ); simpl in *.
642
643     destruct case_nd_id0.      apply nd_id0.
644     destruct case_nd_id1.      apply nd_id1.
645     destruct case_nd_weak.     apply nd_weak.
646     destruct case_nd_copy.     apply nd_copy.
647     destruct case_nd_prod.     apply (nd_prod (nd_map _ _ lpf) (nd_map _ _ rpf)).
648     destruct case_nd_comp.     apply (nd_comp (nd_map _ _ top) (nd_map _ _ bot)).
649     destruct case_nd_cancell.  apply nd_cancell.
650     destruct case_nd_cancelr.  apply nd_cancelr.
651     destruct case_nd_llecnac.  apply nd_llecnac.
652     destruct case_nd_rlecnac.  apply nd_rlecnac.
653     destruct case_nd_assoc.    apply nd_assoc.
654     destruct case_nd_cossa.    apply nd_cossa.
655     apply r. apply rule.
656     Defined.
657
658 (* "map" over natural deduction proofs, where the result proof has different judgments *)
659 Definition nd_map'
660   : forall
661     {Judgment0}{Rule0}{Judgment1}{Rule1}
662     (f:Judgment0->Judgment1)
663     (r:forall h c, Rule0 h c -> @ND Judgment1 Rule1 (mapOptionTree f h) (mapOptionTree f c))
664     {h}{c}
665     (pf:@ND Judgment0 Rule0 h c)
666     ,
667     @ND Judgment1 Rule1 (mapOptionTree f h) (mapOptionTree f c).
668     intros Judgment0 Rule0 Judgment1 Rule1 f r.
669     
670     refine ((fix nd_map' h c pf {struct pf} :=
671      ((match pf
672          in @ND _ _ H C
673           return
674           @ND Judgment1 Rule1 (mapOptionTree f H) (mapOptionTree f C)
675         with
676         | nd_id0                     => let case_nd_id0     := tt in _
677         | nd_id1     h               => let case_nd_id1     := tt in _
678         | nd_weak1   h               => let case_nd_weak    := tt in _
679         | nd_copy    h               => let case_nd_copy    := tt in _
680         | nd_prod    _ _ _ _ lpf rpf => let case_nd_prod    := tt in _
681         | nd_comp    _ _ _   top bot => let case_nd_comp    := tt in _
682         | nd_rule    _ _     rule    => let case_nd_rule    := tt in _
683         | nd_cancell _               => let case_nd_cancell := tt in _
684         | nd_cancelr _               => let case_nd_cancelr := tt in _
685         | nd_llecnac _               => let case_nd_llecnac := tt in _
686         | nd_rlecnac _               => let case_nd_rlecnac := tt in _
687         | nd_assoc   _ _ _           => let case_nd_assoc   := tt in _
688         | nd_cossa   _ _ _           => let case_nd_cossa   := tt in _
689       end))) ); simpl in *.
690
691     destruct case_nd_id0.      apply nd_id0.
692     destruct case_nd_id1.      apply nd_id1.
693     destruct case_nd_weak.     apply nd_weak.
694     destruct case_nd_copy.     apply nd_copy.
695     destruct case_nd_prod.     apply (nd_prod (nd_map' _ _ lpf) (nd_map' _ _ rpf)).
696     destruct case_nd_comp.     apply (nd_comp (nd_map' _ _ top) (nd_map' _ _ bot)).
697     destruct case_nd_cancell.  apply nd_cancell.
698     destruct case_nd_cancelr.  apply nd_cancelr.
699     destruct case_nd_llecnac.  apply nd_llecnac.
700     destruct case_nd_rlecnac.  apply nd_rlecnac.
701     destruct case_nd_assoc.    apply nd_assoc.
702     destruct case_nd_cossa.    apply nd_cossa.
703     apply r. apply rule.
704     Defined.
705
706 (* witnesses the fact that every Rule in a particular proof satisfies the given predicate *)
707 Inductive nd_property {Judgment}{Rule}(P:forall h c, @Rule h c -> Prop) : forall {h}{c}, @ND Judgment Rule h c -> Prop :=
708   | nd_property_structural      : forall h c pf, Structural pf -> @nd_property _ _ P h c pf
709   | nd_property_prod            : forall h0 c0 pf0 h1 c1 pf1,
710     @nd_property _ _ P h0 c0 pf0 -> @nd_property _ _ P h1 c1 pf1 -> @nd_property _ _ P _ _ (nd_prod pf0 pf1)
711   | nd_property_comp            : forall h0 c0 pf0  c1 pf1,
712     @nd_property _ _ P h0 c0 pf0 -> @nd_property _ _ P c0 c1 pf1 -> @nd_property _ _ P _ _ (nd_comp pf0 pf1)
713   | nd_property_rule            : forall h c r, P h c r -> @nd_property _ _ P h c (nd_rule r).
714   Hint Constructors nd_property.
715
716 (* witnesses the fact that every Rule in a particular proof satisfies the given predicate (for ClosedSIND) *)
717 Inductive cnd_property {Judgment}{Rule}(P:forall h c, @Rule h c -> Prop) : forall {c}, @ClosedSIND Judgment Rule c -> Prop :=
718 | cnd_property_weak            : @cnd_property _ _ P _ cnd_weak
719 | cnd_property_rule            : forall h c r cnd',
720   P h c r ->
721   @cnd_property _ _ P h cnd' ->
722   @cnd_property _ _ P c (cnd_rule _ _ cnd' r)
723 | cnd_property_branch          :
724   forall c1 c2 cnd1 cnd2,
725   @cnd_property _ _ P c1 cnd1 ->
726   @cnd_property _ _ P c2 cnd2 ->
727   @cnd_property _ _ P _  (cnd_branch _ _ cnd1 cnd2).
728
729 (* witnesses the fact that every Rule in a particular proof satisfies the given predicate (for SIND) *)
730 Inductive scnd_property {Judgment}{Rule}(P:forall h c, @Rule h c -> Prop) : forall {h c}, @SIND Judgment Rule h c -> Prop :=
731 | scnd_property_weak            : forall c, @scnd_property _ _ P _ _ (scnd_weak c)
732 | scnd_property_comp            : forall h x c r cnd',
733   P x [c] r ->
734   @scnd_property _ _ P h x cnd' ->
735   @scnd_property _ _ P h _ (scnd_comp _ _ _ cnd' r)
736 | scnd_property_branch          :
737   forall x c1 c2 cnd1 cnd2,
738   @scnd_property _ _ P x c1 cnd1 ->
739   @scnd_property _ _ P x c2 cnd2 ->
740   @scnd_property _ _ P x _  (scnd_branch _ _ _ cnd1 cnd2).
741
742 (* renders a proof as LaTeX code *)
743 Section ToLatex.
744
745   Context {Judgment : Type}.
746   Context {Rule     : forall (hypotheses:Tree ??Judgment)(conclusion:Tree ??Judgment), Type}.
747   Context {JudgmentToLatexMath : ToLatexMath Judgment}.
748   Context {RuleToLatexMath     : forall h c, ToLatexMath (Rule h c)}.
749   
750   Open Scope string_scope.
751
752   Definition judgments2latex (j:Tree ??Judgment) := treeToLatexMath (mapOptionTree toLatexMath j).
753
754   Definition eolL : LatexMath := rawLatexMath eol.
755
756   (* invariant: each proof shall emit its hypotheses visibly, except nd_id0 *)
757   Section SIND_toLatex.
758
759     (* indicates which rules should be hidden (omitted) from the rendered proof; useful for structural operations *)
760     Context (hideRule : forall h c (r:Rule h c), bool).
761
762     Fixpoint SIND_toLatexMath {h}{c}(pns:SIND(Rule:=Rule) h c) : LatexMath :=
763       match pns with
764         | scnd_branch ht c1 c2 pns1 pns2 => SIND_toLatexMath pns1 +++ rawLatexMath " \hspace{1cm} " +++ SIND_toLatexMath pns2
765         | scnd_weak     c                => rawLatexMath ""
766         | scnd_comp ht ct c pns rule     => if hideRule _ _ rule
767                                             then SIND_toLatexMath pns
768                                             else rawLatexMath "\trfrac["+++ toLatexMath rule +++ rawLatexMath "]{" +++ eolL +++
769                                               SIND_toLatexMath pns +++ rawLatexMath "}{" +++ eolL +++
770                                               toLatexMath c +++ rawLatexMath "}" +++ eolL
771       end.
772   End SIND_toLatex.
773
774   (* this is a work-in-progress; please use SIND_toLatexMath for now *)
775   Fixpoint nd_toLatexMath {h}{c}(nd:@ND _ Rule h c)(indent:string) :=
776     match nd with
777       | nd_id0                      => rawLatexMath indent +++
778                                        rawLatexMath "% nd_id0 " +++ eolL
779       | nd_id1  h'                  => rawLatexMath indent +++
780                                        rawLatexMath "% nd_id1 "+++ judgments2latex h +++ eolL
781       | nd_weak1 h'                 => rawLatexMath indent +++
782                                        rawLatexMath "\inferrule*[Left=ndWeak]{" +++ toLatexMath h' +++ rawLatexMath "}{ }" +++ eolL
783       | nd_copy h'                  => rawLatexMath indent +++
784                                        rawLatexMath "\inferrule*[Left=ndCopy]{"+++judgments2latex h+++
785                                                          rawLatexMath "}{"+++judgments2latex c+++rawLatexMath "}" +++ eolL
786       | nd_prod h1 h2 c1 c2 pf1 pf2 => rawLatexMath indent +++
787                                        rawLatexMath "% prod " +++ eolL +++
788                                        rawLatexMath indent +++
789                                        rawLatexMath "\begin{array}{c c}" +++ eolL +++
790                                        (nd_toLatexMath pf1 ("  "+++indent)) +++
791                                        rawLatexMath indent +++
792                                        rawLatexMath " & " +++ eolL +++
793                                        (nd_toLatexMath pf2 ("  "+++indent)) +++
794                                        rawLatexMath indent +++
795                                        rawLatexMath "\end{array}"
796       | nd_comp h  m     c  pf1 pf2 => rawLatexMath indent +++
797                                        rawLatexMath "% comp " +++ eolL +++
798                                        rawLatexMath indent +++
799                                        rawLatexMath "\begin{array}{c}" +++ eolL +++
800                                        (nd_toLatexMath pf1 ("  "+++indent)) +++
801                                        rawLatexMath indent +++
802                                        rawLatexMath " \\ " +++ eolL +++
803                                        (nd_toLatexMath pf2 ("  "+++indent)) +++
804                                        rawLatexMath indent +++
805                                        rawLatexMath "\end{array}"
806       | nd_cancell a                => rawLatexMath indent +++
807                                        rawLatexMath "% nd-cancell " +++ (judgments2latex a) +++ eolL
808       | nd_cancelr a                => rawLatexMath indent +++
809                                        rawLatexMath "% nd-cancelr " +++ (judgments2latex a) +++ eolL
810       | nd_llecnac a                => rawLatexMath indent +++
811                                        rawLatexMath "% nd-llecnac " +++ (judgments2latex a) +++ eolL
812       | nd_rlecnac a                => rawLatexMath indent +++
813                                        rawLatexMath "% nd-rlecnac " +++ (judgments2latex a) +++ eolL
814       | nd_assoc   a b c            => rawLatexMath ""
815       | nd_cossa   a b c            => rawLatexMath ""
816       | nd_rule    h c r            => toLatexMath r
817     end.
818
819 End ToLatex.
820
821 Close Scope pf_scope.
822 Close Scope nd_scope.