X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;ds=sidebyside;f=src%2FNaturalDeduction.v;h=407948b134bc763b3cad1963a8ce341bc5df117d;hb=85e4f0fd6b0673c1cc763eeb2585b7dc3d388455;hp=047ab87746125c9747876a87ef004e0774f2b6b2;hpb=112daf37524662d6d2267d3f7e50ff3522683b8f;p=coq-hetmet.git diff --git a/src/NaturalDeduction.v b/src/NaturalDeduction.v index 047ab87..407948b 100644 --- a/src/NaturalDeduction.v +++ b/src/NaturalDeduction.v @@ -1,5 +1,8 @@ (*********************************************************************************************************************************) -(* NaturalDeduction: structurally explicit proofs in Coq *) +(* NaturalDeduction: *) +(* *) +(* Structurally explicit natural deduction proofs. *) +(* *) (*********************************************************************************************************************************) Generalizable All Variables. @@ -218,7 +221,6 @@ Section Natural_Deduction. * emit as LaTeX code. *) Inductive SCND : Tree ??Judgment -> Tree ??Judgment -> Type := - | scnd_axiom : forall c , SCND [] [c] | scnd_comp : forall ht ct c , SCND ht ct -> Rule ct [c] -> SCND ht [c] | scnd_weak : forall c , SCND c [] | scnd_leaf : forall ht c , SCND ht [c] -> SCND ht [c] @@ -226,14 +228,6 @@ Section Natural_Deduction. . Hint Constructors SCND. - Definition mkSCNDAxioms (h:Tree ??Judgment) : SCND [] h. - induction h. - destruct a. - apply scnd_leaf. apply scnd_axiom. - apply scnd_weak. - apply scnd_branch; auto. - Defined. - (* Any ND whose primitive Rules have at most one conclusion (note that nd_prod is allowed!) can be turned into an SCND. *) Definition mkSCND (all_rules_one_conclusion : forall h c1 c2, Rule h (c1,,c2) -> False) : forall h x c, ND x c -> SCND h x -> SCND h c. @@ -265,18 +259,16 @@ Section Natural_Deduction. (* a "ClosedND" is a proof with no open hypotheses and no multi-conclusion rules *) Inductive ClosedND : Tree ??Judgment -> Type := | cnd_weak : ClosedND [] - | cnd_rule : forall h c , ClosedND h -> Rule h c -> ClosedND c + | cnd_rule : forall h c , ClosedND h -> Rule h c -> ClosedND c | cnd_branch : forall c1 c2, ClosedND c1 -> ClosedND c2 -> ClosedND (c1,,c2) . - (* (* we can turn an SCND without hypotheses into a ClosedND *) Definition closedFromSCND h c (pn2:SCND h c)(cnd:ClosedND h) : ClosedND c. refine ((fix closedFromPnodes h c (pn2:SCND h c)(cnd:ClosedND h) {struct pn2} := (match pn2 in SCND H C return H=h -> C=c -> _ with | scnd_weak c => let case_weak := tt in _ | scnd_leaf ht z pn' => let case_leaf := tt in let qq := closedFromPnodes _ _ pn' in _ - | scnd_axiom c => let case_axiom := tt in _ | scnd_comp ht ct c pn' rule => let case_comp := tt in let qq := closedFromPnodes _ _ pn' in _ | scnd_branch ht c1 c2 pn' pn'' => let case_branch := tt in let q1 := closedFromPnodes _ _ pn' in @@ -284,10 +276,6 @@ Section Natural_Deduction. end (refl_equal _) (refl_equal _))) h c pn2 cnd). - destruct case_axiom. - intros; subst. - (* FIXME *) - destruct case_comp. intros. clear pn2. @@ -313,7 +301,14 @@ Section Natural_Deduction. apply q1. subst. apply cnd0. apply q2. subst. apply cnd0. Defined. - *) + + (* undo the above *) + Fixpoint closedNDtoNormalND {c}(cnd:ClosedND c) : ND [] c := + match cnd in ClosedND C return ND [] C with + | cnd_weak => nd_id0 + | cnd_rule h c cndh rhc => closedNDtoNormalND cndh ;; nd_rule rhc + | cnd_branch c1 c2 cnd1 cnd2 => nd_llecnac ;; nd_prod (closedNDtoNormalND cnd1) (closedNDtoNormalND cnd2) + end. Close Scope nd_scope. Open Scope pf_scope.