remove ClosedSIND (use "SIND []" instead)
[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   (* Natural Deduction systems whose judgments happen to be pairs of the same type *)
329   Section SequentND.
330     Context {S:Type}.                   (* type of sequent components *)
331     Context {sequent:S->S->Judgment}.   (* pairing operation which forms a sequent from its halves *)
332     Notation "a |= b" := (sequent a b).
333
334     (* a SequentND is a natural deduction whose judgments are sequents, has initial sequents, and has a cut rule *)
335     Class SequentND :=
336     { snd_initial : forall a,                           [ ] /⋯⋯/ [ a |= a ]
337     ; snd_cut     : forall a b c,  [ a |= b ] ,, [ b |= c ] /⋯⋯/ [ a |= c ]
338     }.
339
340     Context (sequentND:SequentND).
341     Context (ndr:ND_Relation).
342
343     (*
344      * A predicate singling out structural rules, initial sequents,
345      * and cut rules.
346      *
347      * Proofs using only structural rules cannot add or remove
348      * judgments - their hypothesis and conclusion judgment-trees will
349      * differ only in "parenthesization" and the presence/absence of
350      * empty leaves.  This means that a proof involving only
351      * structural rules, cut, and initial sequents can ADD new
352      * non-empty judgment-leaves only via snd_initial, and can only
353      * REMOVE non-empty judgment-leaves only via snd_cut.  Since the
354      * initial sequent is a left and right identity for cut, and cut
355      * is associative, any two proofs (with the same hypotheses and
356      * conclusions) using only structural rules, cut, and initial
357      * sequents are logically indistinguishable - their differences
358      * are logically insignificant.
359      *
360      * Note that it is important that nd_weak and nd_copy aren't
361      * considered to be "structural".
362      *)
363     Inductive SequentND_Inert : forall h c, h/⋯⋯/c -> Prop :=
364     | snd_inert_initial   : forall a,                     SequentND_Inert _ _ (snd_initial a)
365     | snd_inert_cut       : forall a b c,                 SequentND_Inert _ _ (snd_cut a b c)
366     | snd_inert_structural: forall a b f, Structural f -> SequentND_Inert a b f
367     .
368
369     (* An ND_Relation for a sequent deduction should not distinguish between two proofs having the same hypotheses and conclusions
370      * if those proofs use only initial sequents, cut, and structural rules (see comment above) *)
371     Class SequentND_Relation :=
372     { sndr_ndr   := ndr
373     ; sndr_inert : forall a b (f g:a/⋯⋯/b),
374       NDPredicateClosure SequentND_Inert f -> 
375       NDPredicateClosure SequentND_Inert g -> 
376       ndr_eqv f g }.
377
378   End SequentND.
379
380   (* Deductions on sequents whose antecedent is a tree of propositions (i.e. a context) *)
381   Section ContextND.
382     Context {P:Type}{sequent:Tree ??P -> Tree ??P -> Judgment}.
383     Context {snd:SequentND(sequent:=sequent)}.
384     Notation "a |= b" := (sequent a b).
385
386     (* Note that these rules mirror nd_{cancell,cancelr,rlecnac,llecnac,assoc,cossa} but are completely separate from them *)
387     Class ContextND :=
388     { cnd_ant_assoc     : forall x a b c, ND [((a,,b),,c) |= x]     [(a,,(b,,c)) |= x   ]
389     ; cnd_ant_cossa     : forall x a b c, ND [(a,,(b,,c)) |= x]     [((a,,b),,c) |= x   ]
390     ; cnd_ant_cancell   : forall x a    , ND [  [],,a     |= x]     [        a   |= x   ]
391     ; cnd_ant_cancelr   : forall x a    , ND [a,,[]       |= x]     [        a   |= x   ]
392     ; cnd_ant_llecnac   : forall x a    , ND [      a     |= x]     [    [],,a   |= x   ]
393     ; cnd_ant_rlecnac   : forall x a    , ND [      a     |= x]     [    a,,[]   |= x   ]
394     ; cnd_expand_left   : forall a b c  , ND [          a |= b]     [ c,,a       |= c,,b]
395     ; cnd_expand_right  : forall a b c  , ND [          a |= b]     [ a,,c       |= b,,c]
396     ; cnd_snd           := snd
397     }.
398
399     Context `(ContextND).
400
401     (*
402      * A predicate singling out initial sequents, cuts, context expansion,
403      * and structural rules.
404      *
405      * Any two proofs (with the same hypotheses and conclusions) whose
406      * non-structural rules do nothing other than expand contexts,
407      * re-arrange contexts, or introduce additional initial-sequent
408      * conclusions are indistinguishable.  One important consequence
409      * is that asking for a small initial sequent and then expanding
410      * it using cnd_expand_{right,left} is no different from simply
411      * asking for the larger initial sequent in the first place.
412      *
413      *)
414     Inductive ContextND_Inert : forall h c, h/⋯⋯/c -> Prop :=
415     | cnd_inert_initial           : forall a,                     ContextND_Inert _ _ (snd_initial a)
416     | cnd_inert_cut               : forall a b c,                 ContextND_Inert _ _ (snd_cut a b c)
417     | cnd_inert_structural        : forall a b f, Structural f -> ContextND_Inert a b f
418     | cnd_inert_cnd_ant_assoc     : forall x a b c, ContextND_Inert _ _ (cnd_ant_assoc   x a b c)
419     | cnd_inert_cnd_ant_cossa     : forall x a b c, ContextND_Inert _ _ (cnd_ant_cossa   x a b c)
420     | cnd_inert_cnd_ant_cancell   : forall x a    , ContextND_Inert _ _ (cnd_ant_cancell x a)
421     | cnd_inert_cnd_ant_cancelr   : forall x a    , ContextND_Inert _ _ (cnd_ant_cancelr x a)
422     | cnd_inert_cnd_ant_llecnac   : forall x a    , ContextND_Inert _ _ (cnd_ant_llecnac x a)
423     | cnd_inert_cnd_ant_rlecnac   : forall x a    , ContextND_Inert _ _ (cnd_ant_rlecnac x a)
424     | cnd_inert_se_expand_left    : forall t g s  , ContextND_Inert _ _ (@cnd_expand_left  _ t g s)
425     | cnd_inert_se_expand_right   : forall t g s  , ContextND_Inert _ _ (@cnd_expand_right _ t g s).
426
427     Class ContextND_Relation {ndr}{sndr:SequentND_Relation _ ndr} :=
428     { cndr_inert : forall {a}{b}(f g:a/⋯⋯/b),
429                            NDPredicateClosure ContextND_Inert f -> 
430                            NDPredicateClosure ContextND_Inert g -> 
431                            ndr_eqv f g
432     ; cndr_sndr  := sndr
433     }.
434
435     (* a proof is Analytic if it does not use cut *)
436     (*
437     Definition Analytic_Rule : NDPredicate :=
438       fun h c f => forall c, not (snd_cut _ _ c = f).
439     Definition AnalyticND := NDPredicateClosure Analytic_Rule.
440
441     (* a proof system has cut elimination if, for every proof, there is an analytic proof with the same conclusion *)
442     Class CutElimination :=
443     { ce_eliminate : forall h c, h/⋯⋯/c -> h/⋯⋯/c
444     ; ce_analytic  : forall h c f, AnalyticND (ce_eliminate h c f)
445     }.
446
447     (* cut elimination is strong if the analytic proof is furthermore equivalent to the original proof *)
448     Class StrongCutElimination :=
449     { sce_ce       <: CutElimination
450     ; ce_strong     : forall h c f, f === ce_eliminate h c f
451     }.
452     *)
453
454   End ContextND.
455
456   Close Scope nd_scope.
457   Open Scope pf_scope.
458
459 End Natural_Deduction.
460
461 Coercion snd_cut   : SequentND >-> Funclass.
462 Coercion cnd_snd   : ContextND >-> SequentND.
463 Coercion sndr_ndr  : SequentND_Relation >-> ND_Relation.
464 Coercion cndr_sndr : ContextND_Relation >-> SequentND_Relation.
465
466 Implicit Arguments ND [ Judgment ].
467 Hint Constructors Structural.
468 Hint Extern 1 => apply nd_id_structural.
469 Hint Extern 1 => apply ndr_builtfrom_structural.
470 Hint Extern 1 => apply nd_structural_id0.     
471 Hint Extern 1 => apply nd_structural_id1.     
472 Hint Extern 1 => apply nd_structural_cancell. 
473 Hint Extern 1 => apply nd_structural_cancelr. 
474 Hint Extern 1 => apply nd_structural_llecnac. 
475 Hint Extern 1 => apply nd_structural_rlecnac. 
476 Hint Extern 1 => apply nd_structural_assoc.   
477 Hint Extern 1 => apply nd_structural_cossa.   
478 Hint Extern 1 => apply ndpc_p.
479 Hint Extern 1 => apply ndpc_prod.
480 Hint Extern 1 => apply ndpc_comp.
481 Hint Extern 1 => apply builtfrom_refl.
482 Hint Extern 1 => apply builtfrom_prod1.
483 Hint Extern 1 => apply builtfrom_prod2.
484 Hint Extern 1 => apply builtfrom_comp1.
485 Hint Extern 1 => apply builtfrom_comp2.
486 Hint Extern 1 => apply builtfrom_P.
487
488 Hint Extern 1 => apply snd_inert_initial.
489 Hint Extern 1 => apply snd_inert_cut.
490 Hint Extern 1 => apply snd_inert_structural.
491
492 Hint Extern 1 => apply cnd_inert_initial.
493 Hint Extern 1 => apply cnd_inert_cut.
494 Hint Extern 1 => apply cnd_inert_structural.
495 Hint Extern 1 => apply cnd_inert_cnd_ant_assoc.
496 Hint Extern 1 => apply cnd_inert_cnd_ant_cossa.
497 Hint Extern 1 => apply cnd_inert_cnd_ant_cancell.
498 Hint Extern 1 => apply cnd_inert_cnd_ant_cancelr.
499 Hint Extern 1 => apply cnd_inert_cnd_ant_llecnac.
500 Hint Extern 1 => apply cnd_inert_cnd_ant_rlecnac.
501 Hint Extern 1 => apply cnd_inert_se_expand_left.
502 Hint Extern 1 => apply cnd_inert_se_expand_right.
503
504 (* This first notation gets its own scope because it can be confusing when we're working with multiple different kinds
505  * of proofs.  When only one kind of proof is in use, it's quite helpful though. *)
506 Notation "H /⋯⋯/ C" := (@ND _ _ H C)             : pf_scope.
507 Notation "a ;; b"   := (nd_comp a b)             : nd_scope.
508 Notation "a ** b"   := (nd_prod a b)             : nd_scope.
509 Notation "[# a #]"  := (nd_rule a)               : nd_scope.
510 Notation "a === b"  := (@ndr_eqv _ _ _ _ _ a b)  : nd_scope.
511
512 (* enable setoid rewriting *)
513 Open Scope nd_scope.
514 Open Scope pf_scope.
515
516 Add Parametric Relation {jt rt ndr h c} : (h/⋯⋯/c) (@ndr_eqv jt rt ndr h c)
517   reflexivity proved by  (@Equivalence_Reflexive  _ _ (ndr_eqv_equivalence h c))
518   symmetry proved by     (@Equivalence_Symmetric  _ _ (ndr_eqv_equivalence h c))
519   transitivity proved by (@Equivalence_Transitive _ _ (ndr_eqv_equivalence h c))
520     as parametric_relation_ndr_eqv.
521   Add Parametric Morphism {jt rt ndr h x c} : (@nd_comp jt rt h x c)
522   with signature ((ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)))
523     as parametric_morphism_nd_comp.
524     intros; apply ndr_comp_respects; auto.
525     Defined.
526   Add Parametric Morphism {jt rt ndr a b c d} : (@nd_prod jt rt a b c d)
527   with signature ((ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)) ==> (ndr_eqv(ND_Relation:=ndr)))
528     as parametric_morphism_nd_prod.
529     intros; apply ndr_prod_respects; auto.
530     Defined.
531
532 Section ND_Relation_Facts.
533   Context `{ND_Relation}.
534
535   (* useful *)
536   Lemma ndr_comp_right_identity : forall h c (f:h/⋯⋯/c), ndr_eqv (f ;; nd_id c) f.
537     intros; apply (ndr_builtfrom_structural f); auto.
538     Defined.
539
540   (* useful *)
541   Lemma ndr_comp_left_identity : forall h c (f:h/⋯⋯/c), ndr_eqv (nd_id h ;; f) f.
542     intros; apply (ndr_builtfrom_structural f); auto.
543     Defined.
544
545   Ltac nd_prod_preserves_comp_ltac P EQV :=
546     match goal with
547       [ |- context [ (?A ** ?B) ;; (?C ** ?D) ] ] => 
548         set (@ndr_prod_preserves_comp _ _ EQV _ _ A _ _ B _ C _ D) as P
549     end.
550
551   Lemma nd_swap A B C D (f:ND _ A B) (g:ND _ C D) :
552     (f ** nd_id C) ;; (nd_id B ** g) ===
553     (nd_id A ** g) ;; (f ** nd_id D).
554     setoid_rewrite <- ndr_prod_preserves_comp.
555     setoid_rewrite ndr_comp_left_identity.
556     setoid_rewrite ndr_comp_right_identity.
557     reflexivity.
558     Qed.
559
560   (* this tactical searches the environment; setoid_rewrite doesn't seem to be able to do that properly sometimes *)
561   Ltac nd_swap_ltac P EQV :=
562     match goal with
563       [ |- context [ (?F ** nd_id _) ;; (nd_id _ ** ?G) ] ] => 
564         set (@nd_swap _ _ EQV _ _ _ _ F G) as P
565     end.
566
567   Lemma nd_prod_split_left A B C D (f:ND _ A B) (g:ND _ B C) :
568     nd_id D ** (f ;; g) ===
569     (nd_id D ** f) ;; (nd_id D ** g).
570     setoid_rewrite <- ndr_prod_preserves_comp.
571     setoid_rewrite ndr_comp_left_identity.
572     reflexivity.
573     Qed.
574
575   Lemma nd_prod_split_right A B C D (f:ND _ A B) (g:ND _ B C) :
576     (f ;; g) ** nd_id D ===
577     (f ** nd_id D) ;; (g ** nd_id D).
578     setoid_rewrite <- ndr_prod_preserves_comp.
579     setoid_rewrite ndr_comp_left_identity.
580     reflexivity.
581     Qed.
582
583 End ND_Relation_Facts.
584
585 (* a generalization of the procedure used to build (nd_id n) from nd_id0 and nd_id1 *)
586 Definition nd_replicate
587   : forall
588     {Judgment}{Rule}{Ob}
589     (h:Ob->Judgment)
590     (c:Ob->Judgment)
591     (j:Tree ??Ob),
592     (forall (o:Ob), @ND Judgment Rule [h o] [c o]) ->
593     @ND Judgment Rule (mapOptionTree h j) (mapOptionTree c j).
594   intros.
595   induction j; simpl.
596     destruct a; simpl.
597     apply X.
598     apply nd_id0.
599     apply nd_prod; auto.
600     Defined.
601
602 (* "map" over natural deduction proofs, where the result proof has the same judgments (but different rules) *)
603 Definition nd_map
604   : forall
605     {Judgment}{Rule0}{Rule1}
606     (r:forall h c, Rule0 h c -> @ND Judgment Rule1 h c)
607     {h}{c}
608     (pf:@ND Judgment Rule0 h c)
609     ,
610     @ND Judgment Rule1 h c.
611     intros Judgment Rule0 Rule1 r.
612
613     refine ((fix nd_map h c pf {struct pf} :=
614      ((match pf
615          in @ND _ _ H C
616           return
617           @ND Judgment Rule1 H C
618         with
619         | nd_id0                     => let case_nd_id0     := tt in _
620         | nd_id1     h               => let case_nd_id1     := tt in _
621         | nd_weak1   h               => let case_nd_weak    := tt in _
622         | nd_copy    h               => let case_nd_copy    := tt in _
623         | nd_prod    _ _ _ _ lpf rpf => let case_nd_prod    := tt in _
624         | nd_comp    _ _ _   top bot => let case_nd_comp    := tt in _
625         | nd_rule    _ _     rule    => let case_nd_rule    := tt in _
626         | nd_cancell _               => let case_nd_cancell := tt in _
627         | nd_cancelr _               => let case_nd_cancelr := tt in _
628         | nd_llecnac _               => let case_nd_llecnac := tt in _
629         | nd_rlecnac _               => let case_nd_rlecnac := tt in _
630         | nd_assoc   _ _ _           => let case_nd_assoc   := tt in _
631         | nd_cossa   _ _ _           => let case_nd_cossa   := tt in _
632       end))) ); simpl in *.
633
634     destruct case_nd_id0.      apply nd_id0.
635     destruct case_nd_id1.      apply nd_id1.
636     destruct case_nd_weak.     apply nd_weak.
637     destruct case_nd_copy.     apply nd_copy.
638     destruct case_nd_prod.     apply (nd_prod (nd_map _ _ lpf) (nd_map _ _ rpf)).
639     destruct case_nd_comp.     apply (nd_comp (nd_map _ _ top) (nd_map _ _ bot)).
640     destruct case_nd_cancell.  apply nd_cancell.
641     destruct case_nd_cancelr.  apply nd_cancelr.
642     destruct case_nd_llecnac.  apply nd_llecnac.
643     destruct case_nd_rlecnac.  apply nd_rlecnac.
644     destruct case_nd_assoc.    apply nd_assoc.
645     destruct case_nd_cossa.    apply nd_cossa.
646     apply r. apply rule.
647     Defined.
648
649 (* "map" over natural deduction proofs, where the result proof has different judgments *)
650 Definition nd_map'
651   : forall
652     {Judgment0}{Rule0}{Judgment1}{Rule1}
653     (f:Judgment0->Judgment1)
654     (r:forall h c, Rule0 h c -> @ND Judgment1 Rule1 (mapOptionTree f h) (mapOptionTree f c))
655     {h}{c}
656     (pf:@ND Judgment0 Rule0 h c)
657     ,
658     @ND Judgment1 Rule1 (mapOptionTree f h) (mapOptionTree f c).
659     intros Judgment0 Rule0 Judgment1 Rule1 f r.
660     
661     refine ((fix nd_map' h c pf {struct pf} :=
662      ((match pf
663          in @ND _ _ H C
664           return
665           @ND Judgment1 Rule1 (mapOptionTree f H) (mapOptionTree f C)
666         with
667         | nd_id0                     => let case_nd_id0     := tt in _
668         | nd_id1     h               => let case_nd_id1     := tt in _
669         | nd_weak1   h               => let case_nd_weak    := tt in _
670         | nd_copy    h               => let case_nd_copy    := tt in _
671         | nd_prod    _ _ _ _ lpf rpf => let case_nd_prod    := tt in _
672         | nd_comp    _ _ _   top bot => let case_nd_comp    := tt in _
673         | nd_rule    _ _     rule    => let case_nd_rule    := tt in _
674         | nd_cancell _               => let case_nd_cancell := tt in _
675         | nd_cancelr _               => let case_nd_cancelr := tt in _
676         | nd_llecnac _               => let case_nd_llecnac := tt in _
677         | nd_rlecnac _               => let case_nd_rlecnac := tt in _
678         | nd_assoc   _ _ _           => let case_nd_assoc   := tt in _
679         | nd_cossa   _ _ _           => let case_nd_cossa   := tt in _
680       end))) ); simpl in *.
681
682     destruct case_nd_id0.      apply nd_id0.
683     destruct case_nd_id1.      apply nd_id1.
684     destruct case_nd_weak.     apply nd_weak.
685     destruct case_nd_copy.     apply nd_copy.
686     destruct case_nd_prod.     apply (nd_prod (nd_map' _ _ lpf) (nd_map' _ _ rpf)).
687     destruct case_nd_comp.     apply (nd_comp (nd_map' _ _ top) (nd_map' _ _ bot)).
688     destruct case_nd_cancell.  apply nd_cancell.
689     destruct case_nd_cancelr.  apply nd_cancelr.
690     destruct case_nd_llecnac.  apply nd_llecnac.
691     destruct case_nd_rlecnac.  apply nd_rlecnac.
692     destruct case_nd_assoc.    apply nd_assoc.
693     destruct case_nd_cossa.    apply nd_cossa.
694     apply r. apply rule.
695     Defined.
696
697 (* witnesses the fact that every Rule in a particular proof satisfies the given predicate *)
698 Inductive nd_property {Judgment}{Rule}(P:forall h c, @Rule h c -> Prop) : forall {h}{c}, @ND Judgment Rule h c -> Prop :=
699   | nd_property_structural      : forall h c pf, Structural pf -> @nd_property _ _ P h c pf
700   | nd_property_prod            : forall h0 c0 pf0 h1 c1 pf1,
701     @nd_property _ _ P h0 c0 pf0 -> @nd_property _ _ P h1 c1 pf1 -> @nd_property _ _ P _ _ (nd_prod pf0 pf1)
702   | nd_property_comp            : forall h0 c0 pf0  c1 pf1,
703     @nd_property _ _ P h0 c0 pf0 -> @nd_property _ _ P c0 c1 pf1 -> @nd_property _ _ P _ _ (nd_comp pf0 pf1)
704   | nd_property_rule            : forall h c r, P h c r -> @nd_property _ _ P h c (nd_rule r).
705   Hint Constructors nd_property.
706
707 (* witnesses the fact that every Rule in a particular proof satisfies the given predicate (for SIND) *)
708 Inductive scnd_property {Judgment}{Rule}(P:forall h c, @Rule h c -> Prop) : forall {h c}, @SIND Judgment Rule h c -> Prop :=
709 | scnd_property_weak            : forall c, @scnd_property _ _ P _ _ (scnd_weak c)
710 | scnd_property_comp            : forall h x c r cnd',
711   P x [c] r ->
712   @scnd_property _ _ P h x cnd' ->
713   @scnd_property _ _ P h _ (scnd_comp _ _ _ cnd' r)
714 | scnd_property_branch          :
715   forall x c1 c2 cnd1 cnd2,
716   @scnd_property _ _ P x c1 cnd1 ->
717   @scnd_property _ _ P x c2 cnd2 ->
718   @scnd_property _ _ P x _  (scnd_branch _ _ _ cnd1 cnd2).
719
720 (* renders a proof as LaTeX code *)
721 Section ToLatex.
722
723   Context {Judgment : Type}.
724   Context {Rule     : forall (hypotheses:Tree ??Judgment)(conclusion:Tree ??Judgment), Type}.
725   Context {JudgmentToLatexMath : ToLatexMath Judgment}.
726   Context {RuleToLatexMath     : forall h c, ToLatexMath (Rule h c)}.
727   
728   Open Scope string_scope.
729
730   Definition judgments2latex (j:Tree ??Judgment) := treeToLatexMath (mapOptionTree toLatexMath j).
731
732   Definition eolL : LatexMath := rawLatexMath eol.
733
734   (* invariant: each proof shall emit its hypotheses visibly, except nd_id0 *)
735   Section SIND_toLatex.
736
737     (* indicates which rules should be hidden (omitted) from the rendered proof; useful for structural operations *)
738     Context (hideRule : forall h c (r:Rule h c), bool).
739
740     Fixpoint SIND_toLatexMath {h}{c}(pns:SIND(Rule:=Rule) h c) : LatexMath :=
741       match pns with
742         | scnd_branch ht c1 c2 pns1 pns2 => SIND_toLatexMath pns1 +++ rawLatexMath " \hspace{1cm} " +++ SIND_toLatexMath pns2
743         | scnd_weak     c                => rawLatexMath ""
744         | scnd_comp ht ct c pns rule     => if hideRule _ _ rule
745                                             then SIND_toLatexMath pns
746                                             else rawLatexMath "\trfrac["+++ toLatexMath rule +++ rawLatexMath "]{" +++ eolL +++
747                                               SIND_toLatexMath pns +++ rawLatexMath "}{" +++ eolL +++
748                                               toLatexMath c +++ rawLatexMath "}" +++ eolL
749       end.
750   End SIND_toLatex.
751
752   (* this is a work-in-progress; please use SIND_toLatexMath for now *)
753   Fixpoint nd_toLatexMath {h}{c}(nd:@ND _ Rule h c)(indent:string) :=
754     match nd with
755       | nd_id0                      => rawLatexMath indent +++
756                                        rawLatexMath "% nd_id0 " +++ eolL
757       | nd_id1  h'                  => rawLatexMath indent +++
758                                        rawLatexMath "% nd_id1 "+++ judgments2latex h +++ eolL
759       | nd_weak1 h'                 => rawLatexMath indent +++
760                                        rawLatexMath "\inferrule*[Left=ndWeak]{" +++ toLatexMath h' +++ rawLatexMath "}{ }" +++ eolL
761       | nd_copy h'                  => rawLatexMath indent +++
762                                        rawLatexMath "\inferrule*[Left=ndCopy]{"+++judgments2latex h+++
763                                                          rawLatexMath "}{"+++judgments2latex c+++rawLatexMath "}" +++ eolL
764       | nd_prod h1 h2 c1 c2 pf1 pf2 => rawLatexMath indent +++
765                                        rawLatexMath "% prod " +++ eolL +++
766                                        rawLatexMath indent +++
767                                        rawLatexMath "\begin{array}{c c}" +++ eolL +++
768                                        (nd_toLatexMath pf1 ("  "+++indent)) +++
769                                        rawLatexMath indent +++
770                                        rawLatexMath " & " +++ eolL +++
771                                        (nd_toLatexMath pf2 ("  "+++indent)) +++
772                                        rawLatexMath indent +++
773                                        rawLatexMath "\end{array}"
774       | nd_comp h  m     c  pf1 pf2 => rawLatexMath indent +++
775                                        rawLatexMath "% comp " +++ eolL +++
776                                        rawLatexMath indent +++
777                                        rawLatexMath "\begin{array}{c}" +++ eolL +++
778                                        (nd_toLatexMath pf1 ("  "+++indent)) +++
779                                        rawLatexMath indent +++
780                                        rawLatexMath " \\ " +++ eolL +++
781                                        (nd_toLatexMath pf2 ("  "+++indent)) +++
782                                        rawLatexMath indent +++
783                                        rawLatexMath "\end{array}"
784       | nd_cancell a                => rawLatexMath indent +++
785                                        rawLatexMath "% nd-cancell " +++ (judgments2latex a) +++ eolL
786       | nd_cancelr a                => rawLatexMath indent +++
787                                        rawLatexMath "% nd-cancelr " +++ (judgments2latex a) +++ eolL
788       | nd_llecnac a                => rawLatexMath indent +++
789                                        rawLatexMath "% nd-llecnac " +++ (judgments2latex a) +++ eolL
790       | nd_rlecnac a                => rawLatexMath indent +++
791                                        rawLatexMath "% nd-rlecnac " +++ (judgments2latex a) +++ eolL
792       | nd_assoc   a b c            => rawLatexMath ""
793       | nd_cossa   a b c            => rawLatexMath ""
794       | nd_rule    h c r            => toLatexMath r
795     end.
796
797 End ToLatex.
798
799 Close Scope pf_scope.
800 Close Scope nd_scope.