X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskProofToStrong.v;h=b6e8efee00b3730606325935a866162a67bb9ffa;hp=06f97a1202d3b02d7e219afc2abab378260f0775;hb=93ac0d63048027161f816c451a7954fb8a6470c0;hpb=64d416692bda1d36c33b5efa245d46dcf546ad4a diff --git a/src/HaskProofToStrong.v b/src/HaskProofToStrong.v index 06f97a1..b6e8efe 100644 --- a/src/HaskProofToStrong.v +++ b/src/HaskProofToStrong.v @@ -239,6 +239,7 @@ Section HaskProofToStrong. with | RLeft h c ctx r => let case_RLeft := tt in (fun e => _) (urule2expr _ _ _ r) | RRight h c ctx r => let case_RRight := tt in (fun e => _) (urule2expr _ _ _ r) + | RId a => let case_RId := tt in _ | RCanL a => let case_RCanL := tt in _ | RCanR a => let case_RCanR := tt in _ | RuCanL a => let case_RuCanL := tt in _ @@ -251,6 +252,9 @@ Section HaskProofToStrong. | RComp a b c f g => let case_RComp := tt in (fun e1 e2 => _) (urule2expr _ _ _ f) (urule2expr _ _ _ g) end); clear urule2expr; intros. + destruct case_RId. + apply X. + destruct case_RCanL. simpl; unfold ujudg2exprType; intros. simpl in X. @@ -569,7 +573,6 @@ Section HaskProofToStrong. destruct case_RGlobal. apply ILeaf; simpl; intros; refine (return ILeaf _ _). apply EGlobal. - apply wev. destruct case_RLam. apply ILeaf. @@ -633,7 +636,7 @@ Section HaskProofToStrong. apply ileaf in q1'. apply ileaf in q2'. simpl in *. - apply (EApp _ _ _ _ _ _ q1' q2'). + apply (EApp _ _ _ _ _ _ q2' q1'). destruct case_RLet. apply ILeaf. @@ -715,6 +718,7 @@ Section HaskProofToStrong. inversion X; subst; clear X. apply (@ELetRec _ _ _ _ _ _ _ varstypes). + auto. apply (@letrec_helper Γ Δ t varstypes). rewrite <- pf2 in X1. rewrite mapOptionTree_compose. @@ -758,15 +762,12 @@ Section HaskProofToStrong. apply H2. Defined. - Definition closed2expr : forall c (pn:@ClosedSIND _ Rule c), ITree _ judg2exprType c. - refine (( - fix closed2expr' j (pn:@ClosedSIND _ Rule j) {struct pn} : ITree _ judg2exprType j := - match pn in @ClosedSIND _ _ J return ITree _ judg2exprType J with - | cnd_weak => let case_nil := tt in INone _ _ - | cnd_rule h c cnd' r => let case_rule := tt in rule2expr _ _ r (closed2expr' _ cnd') - | cnd_branch _ _ c1 c2 => let case_branch := tt in IBranch _ _ (closed2expr' _ c1) (closed2expr' _ c2) - end)); clear closed2expr'; intros; subst. - Defined. + Fixpoint closed2expr h j (pn:@SIND _ Rule h j) {struct pn} : ITree _ judg2exprType h -> ITree _ judg2exprType j := + match pn in @SIND _ _ H J return ITree _ judg2exprType H -> ITree _ judg2exprType J with + | scnd_weak _ => let case_nil := tt in fun _ => INone _ _ + | scnd_comp x h c cnd' r => let case_rule := tt in fun q => rule2expr _ _ r (closed2expr _ _ cnd' q) + | scnd_branch _ _ _ c1 c2 => let case_branch := tt in fun q => IBranch _ _ (closed2expr _ _ c1 q) (closed2expr _ _ c2 q) + end. Lemma manyFresh : forall Γ Σ (ξ0:VV -> LeveledHaskType Γ ★), FreshM { vars : _ & { ξ : VV -> LeveledHaskType Γ ★ & Σ = mapOptionTree ξ vars } }. @@ -804,7 +805,7 @@ Section HaskProofToStrong. {zz:ToString VV} : ND Rule [] [Γ > Δ > Σ |- [τ]] -> FreshM (???{ ξ : _ & Expr Γ Δ ξ τ}). intro pf. - set (closedFromSIND _ _ (mkSIND systemfc_all_rules_one_conclusion _ _ _ pf (scnd_weak [])) cnd_weak) as cnd. + set (mkSIND systemfc_all_rules_one_conclusion _ _ _ pf (scnd_weak [])) as cnd. apply closed2expr in cnd. apply ileaf in cnd. simpl in *. @@ -819,6 +820,7 @@ Section HaskProofToStrong. refine (return OK _). exists ξ. apply (ileaf it). + apply INone. Defined. End HaskProofToStrong.