X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskStrongToProof.v;h=8ae0d0920b8ca99e6f9c301bdb5075667dc2759a;hp=8aa40051950869620b76f6e320f504ca2b547264;hb=3161a8a65cb0190e83d32bde613c3b64dfe31739;hpb=db8c9d54c285980e162e393efd1b7316887e5b80 diff --git a/src/HaskStrongToProof.v b/src/HaskStrongToProof.v index 8aa4005..8ae0d09 100644 --- a/src/HaskStrongToProof.v +++ b/src/HaskStrongToProof.v @@ -549,23 +549,6 @@ match elrb with | ELR_branch Γ Δ ξ lev t1 t2 b1 b2 => (eLetRecContext b1),,(eLetRecContext b2) end. -Definition mkProofCaseBranch {Γ}{Δ}{ξ}{l}{tc}{tbranches}{atypes} -(alt : { sac : _ & { scb : StrongCaseBranchWithVVs _ _ tc atypes sac - & Expr (sac_gamma sac Γ) - (sac_delta sac Γ atypes (weakCK'' Δ)) - (scbwv_xi scb ξ l) - (weakT' tbranches) (weakL' l) } }) - : { sac : _ & ProofCaseBranch tc Γ Δ l tbranches atypes sac }. - destruct alt. - exists x. - exact - {| pcb_freevars := mapOptionTree ξ - (stripOutVars (vec2list (scbwv_exprvars (projT1 s))) - (expr2antecedent (projT2 s))) - |}. - Defined. - - Fixpoint eLetRecTypes {Γ}{Δ}{ξ}{lev}{τ}(elrb:ELetRecBindings Γ Δ ξ lev τ) : Tree ??(HaskType Γ ★) := match elrb with | ELR_nil Γ Δ ξ lev => [] @@ -621,16 +604,16 @@ Definition factorContextLeft (* where the leaf is v *) apply inr. subst. - apply RuCanR. + apply AuCanR. (* where the leaf is NOT v *) apply inl. - apply RuCanL. + apply AuCanL. (* empty leaf *) destruct case_None. apply inl; simpl in *. - apply RuCanR. + apply AuCanR. (* branch *) refine ( @@ -650,45 +633,45 @@ Definition factorContextLeft destruct case_Neither. apply inl. simpl. - eapply RComp; [idtac | apply RuCanL ]. - exact (RComp + eapply AComp; [idtac | apply AuCanL ]. + exact (AComp (* order will not matter because these are central as morphisms *) - (RRight _ (RComp lpf (RCanL _))) - (RLeft _ (RComp rpf (RCanL _)))). + (ARight _ (AComp lpf (ACanL _))) + (ALeft _ (AComp rpf (ACanL _)))). destruct case_Right. apply inr. unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *. fold (stripOutVars (v::nil)). - eapply RComp; [ idtac | eapply pivotContext' ]. - eapply RComp. - eapply RRight. - eapply RComp. + eapply AComp; [ idtac | eapply pivotContext' ]. + eapply AComp. + eapply ARight. + eapply AComp. apply lpf. - apply RCanL. - eapply RLeft. + apply ACanL. + eapply ALeft. apply rpf. destruct case_Left. apply inr. fold (stripOutVars (v::nil)). simpl. - eapply RComp. - eapply RLeft. - eapply RComp. + eapply AComp. + eapply ALeft. + eapply AComp. apply rpf. simpl. - eapply RCanL. - eapply RComp; [ idtac | eapply RCossa ]. - eapply RRight. + eapply ACanL. + eapply AComp; [ idtac | eapply AuAssoc ]. + eapply ARight. apply lpf. destruct case_Both. apply inr. simpl. - eapply RComp; [ idtac | eapply RRight; eapply RCont ]. - eapply RComp; [ eapply RRight; eapply lpf | idtac ]. - eapply RComp; [ eapply RLeft; eapply rpf | idtac ]. + eapply AComp; [ idtac | eapply ARight; eapply ACont ]. + eapply AComp; [ eapply ARight; eapply lpf | idtac ]. + eapply AComp; [ eapply ALeft; eapply rpf | idtac ]. clear lpf rpf. simpl. apply arrangeSwapMiddle. @@ -726,16 +709,16 @@ Definition factorContextRight (* where the leaf is v *) apply inr. subst. - apply RuCanL. + apply AuCanL. (* where the leaf is NOT v *) apply inl. - apply RuCanR. + apply AuCanR. (* empty leaf *) destruct case_None. apply inl; simpl in *. - apply RuCanR. + apply AuCanR. (* branch *) refine ( @@ -754,51 +737,51 @@ Definition factorContextRight destruct case_Neither. apply inl. - eapply RComp; [idtac | apply RuCanR ]. - exact (RComp + eapply AComp; [idtac | apply AuCanR ]. + exact (AComp (* order will not matter because these are central as morphisms *) - (RRight _ (RComp lpf (RCanR _))) - (RLeft _ (RComp rpf (RCanR _)))). + (ARight _ (AComp lpf (ACanR _))) + (ALeft _ (AComp rpf (ACanR _)))). destruct case_Right. apply inr. fold (stripOutVars (v::nil)). - set (RRight (mapOptionTree ξ ctx2) (RComp lpf ((RCanR _)))) as q. + set (ARight (mapOptionTree ξ ctx2) (AComp lpf ((ACanR _)))) as q. simpl in *. - eapply RComp. + eapply AComp. apply q. clear q. clear lpf. unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *. - eapply RComp; [ idtac | apply RAssoc ]. - apply RLeft. + eapply AComp; [ idtac | apply AAssoc ]. + apply ALeft. apply rpf. destruct case_Left. apply inr. unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *. fold (stripOutVars (v::nil)). - eapply RComp; [ idtac | eapply pivotContext ]. - set (RComp rpf (RCanR _ )) as rpf'. - set (RLeft ((mapOptionTree ξ (stripOutVars (v :: nil) ctx1),, [ξ v])) rpf') as qq. + eapply AComp; [ idtac | eapply pivotContext ]. + set (AComp rpf (ACanR _ )) as rpf'. + set (ALeft ((mapOptionTree ξ (stripOutVars (v :: nil) ctx1),, [ξ v])) rpf') as qq. simpl in *. - eapply RComp; [ idtac | apply qq ]. + eapply AComp; [ idtac | apply qq ]. clear qq rpf' rpf. - apply (RRight (mapOptionTree ξ ctx2)). + apply (ARight (mapOptionTree ξ ctx2)). apply lpf. destruct case_Both. apply inr. unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *. fold (stripOutVars (v::nil)). - eapply RComp; [ idtac | eapply copyAndPivotContext ]. + eapply AComp; [ idtac | eapply copyAndPivotContext ]. (* order will not matter because these are central as morphisms *) - exact (RComp (RRight _ lpf) (RLeft _ rpf)). + exact (AComp (ARight _ lpf) (ALeft _ rpf)). Defined. -(* same as before, but use RWeak if necessary *) +(* same as before, but use AWeak if necessary *) Definition factorContextLeftAndWeaken (Γ:TypeEnv)(Δ:CoercionEnv Γ) v (* variable to be pivoted, if found *) @@ -809,8 +792,8 @@ Definition factorContextLeftAndWeaken (mapOptionTree ξ ([v],,(stripOutVars (v::nil) ctx)) ). set (factorContextLeft Γ Δ v ctx ξ) as q. destruct q; auto. - eapply RComp; [ apply a | idtac ]. - refine (RRight _ (RWeak _)). + eapply AComp; [ apply a | idtac ]. + refine (ARight _ (AWeak _)). Defined. Definition factorContextLeftAndWeaken'' @@ -836,7 +819,7 @@ Definition factorContextLeftAndWeaken'' unfold mapOptionTree; simpl in *. intros. rewrite (@stripping_nothing_is_inert Γ); auto. - apply RuCanL. + apply AuCanL. intros. unfold mapOptionTree in *. simpl in *. @@ -855,15 +838,15 @@ Definition factorContextLeftAndWeaken'' unfold stripOutVars in q. rewrite q in IHv1'. clear q. - eapply RComp; [ idtac | eapply RAssoc ]. - eapply RComp; [ idtac | eapply IHv1' ]. + eapply AComp; [ idtac | eapply AAssoc ]. + eapply AComp; [ idtac | eapply IHv1' ]. clear IHv1'. apply IHv2; auto. auto. auto. Defined. -(* same as before, but use RWeak if necessary *) +(* same as before, but use AWeak if necessary *) Definition factorContextRightAndWeaken (Γ:TypeEnv)(Δ:CoercionEnv Γ) v (* variable to be pivoted, if found *) @@ -874,8 +857,8 @@ Definition factorContextRightAndWeaken (mapOptionTree ξ ((stripOutVars (v::nil) ctx),,[v]) ). set (factorContextRight Γ Δ v ctx ξ) as q. destruct q; auto. - eapply RComp; [ apply a | idtac ]. - refine (RLeft _ (RWeak _)). + eapply AComp; [ apply a | idtac ]. + refine (ALeft _ (AWeak _)). Defined. Definition factorContextRightAndWeaken'' @@ -899,7 +882,7 @@ Definition factorContextRightAndWeaken'' unfold mapOptionTree; simpl in *. intros. rewrite (@stripping_nothing_is_inert Γ); auto. - apply RuCanR. + apply AuCanR. intros. unfold mapOptionTree in *. simpl in *. @@ -918,14 +901,14 @@ Definition factorContextRightAndWeaken'' fold X in IHv2'. set (distinct_app _ _ _ H) as H'. destruct H' as [H1 H2]. - set (RComp (IHv1 _ H1) (IHv2' H2)) as qq. - eapply RComp. + set (AComp (IHv1 _ H1) (IHv2' H2)) as qq. + eapply AComp. apply qq. clear qq IHv2' IHv2 IHv1. rewrite strip_swap_lemma. rewrite strip_twice_lemma. rewrite (notin_strip_inert' v1 (leaves v2)). - apply RCossa. + apply AuAssoc. apply distinct_swap. auto. Defined. @@ -961,10 +944,13 @@ Lemma letRecSubproofsToND Γ Δ ξ lev tree branches : destruct q. simpl in *. apply n. - eapply nd_comp; [ idtac | eapply nd_rule; apply RJoin ]. - eapply nd_comp; [ apply nd_llecnac | idtac ]. - apply nd_prod; auto. - Defined. + eapply nd_comp; [ idtac | eapply RCut' ]. + eapply nd_comp; [ apply nd_llecnac | idtac ]. + apply nd_prod. + apply IHX1. + eapply nd_comp; [ idtac | eapply nd_rule; eapply RLeft ]. + apply IHX2. + Defined. Lemma letRecSubproofsToND' Γ Δ ξ lev τ tree : forall branches body (dist:distinct (leaves (mapOptionTree (@fst _ _) tree))), @@ -994,7 +980,7 @@ Lemma letRecSubproofsToND' Γ Δ ξ lev τ tree : eapply nd_comp; [ idtac | eapply nd_rule; apply z ]. clear z. - set (@factorContextRightAndWeaken'' Γ Δ pctx ξ' (eLetRecContext branches,,expr2antecedent body)) as q'. + set (@factorContextLeftAndWeaken'' Γ Δ pctx ξ' (eLetRecContext branches,,expr2antecedent body)) as q'. unfold passback in *; clear passback. unfold pctx in *; clear pctx. set (q' disti) as q''. @@ -1015,10 +1001,14 @@ Lemma letRecSubproofsToND' Γ Δ ξ lev τ tree : simpl. set (letRecSubproofsToND _ _ _ _ _ branches lrsp) as q. - eapply nd_comp; [ idtac | eapply nd_rule; apply RJoin ]. - eapply nd_comp; [ apply nd_rlecnac | idtac ]. - apply nd_prod; auto. - Defined. + + eapply nd_comp; [ idtac | eapply RCut' ]. + eapply nd_comp; [ apply nd_llecnac | idtac ]. + apply nd_prod. + apply q. + eapply nd_comp; [ idtac | eapply nd_rule; eapply RLeft ]. + apply pf. + Defined. Lemma scbwv_coherent {tc}{Γ}{atypes:IList _ (HaskType Γ) _}{sac} : forall scb:StrongCaseBranchWithVVs _ _ tc atypes sac, @@ -1050,6 +1040,21 @@ Lemma scbwv_coherent {tc}{Γ}{atypes:IList _ (HaskType Γ) _}{sac} : Qed. +Definition mkProofCaseBranch {Γ}{Δ}{ξ}{l}{tc}{tbranches}{atypes} +(alt : { sac : _ & { scb : StrongCaseBranchWithVVs _ _ tc atypes sac + & Expr (sac_gamma sac Γ) + (sac_delta sac Γ atypes (weakCK'' Δ)) + (scbwv_xi scb ξ l) + (weakT' tbranches) (weakL' l) } }) + : @StrongAltCon tc * Tree ??(LeveledHaskType Γ ★). + destruct alt. + split. + apply x. + apply (mapOptionTree ξ + (stripOutVars (vec2list (scbwv_exprvars (projT1 s))) + (expr2antecedent (projT2 s)))). + Defined. + Lemma case_lemma : forall Γ Δ ξ l tc tbranches atypes e (alts':Tree ??{sac : StrongAltCon & @@ -1057,7 +1062,7 @@ Lemma case_lemma : forall Γ Δ ξ l tc tbranches atypes e Expr (sac_gamma sac Γ) (sac_delta sac Γ atypes (weakCK'' Δ)) (scbwv_xi scb ξ l) (weakT' tbranches) (weakL' l)}}), - (mapOptionTreeAndFlatten (fun x => pcb_freevars (projT2 x)) + (mapOptionTreeAndFlatten (fun x => snd x) (mapOptionTree mkProofCaseBranch alts')) ,, mapOptionTree ξ (expr2antecedent e) = @@ -1124,13 +1129,15 @@ Definition expr2proof : (sac_delta sac Γ atypes (weakCK'' Δ)) (scbwv_xi scb ξ l) (weakT' tbranches) (weakL' l) } }) - : ND Rule [] (mapOptionTree (fun x => pcb_judg (projT2 (mkProofCaseBranch x))) alts) := + : ND Rule [] (mapOptionTree (fun x => pcb_judg (snd (mkProofCaseBranch x))) alts) := match alts as ALTS return ND Rule [] - (mapOptionTree (fun x => pcb_judg (projT2 (mkProofCaseBranch x))) ALTS) with + (mapOptionTree (fun x => pcb_judg (snd (mkProofCaseBranch x))) ALTS) with | T_Leaf None => let case_nil := tt in _ | T_Branch b1 b2 => let case_branch := tt in (fun b1' b2' => _) (mkdcsp b1) (mkdcsp b2) | T_Leaf (Some x) => - match x as X return ND Rule [] [pcb_judg (projT2 (mkProofCaseBranch X))] with + match x as X return ND Rule [] [@pcb_judg tc Γ Δ l tbranches atypes + (fst (mkProofCaseBranch X)) + (snd (mkProofCaseBranch X))] with existT sac (existT scbx ex) => (fun e' => let case_leaf := tt in _) (expr2proof _ _ _ _ _ ex) end @@ -1182,7 +1189,7 @@ Definition expr2proof : inversion H. destruct case_ELet; intros; simpl in *. - eapply nd_comp; [ idtac | eapply nd_rule; eapply RLet ]. + eapply nd_comp; [ idtac | eapply RLet ]. eapply nd_comp; [ apply nd_rlecnac | idtac ]. apply nd_prod. apply pf_let. @@ -1243,10 +1250,10 @@ Definition expr2proof : destruct case_leaf. clear o x alts alts' e. - eapply nd_comp; [ apply e' | idtac ]. + simpl. + apply (fun x => nd_comp e' x). clear e'. - apply nd_rule. - apply RArrange. + unfold pcb_judg. simpl. rewrite mapleaves'. simpl. @@ -1255,26 +1262,40 @@ Definition expr2proof : rewrite <- mapleaves'. rewrite vec2list_map_list2vec. unfold sac_gamma. - rewrite <- (scbwv_coherent scbx l ξ). rewrite <- vec2list_map_list2vec. rewrite mapleaves'. - set (@factorContextRightAndWeaken'') as q. - unfold scbwv_xi. set (@updating_stripped_tree_is_inert' _ (weakL' l) (weakLT' ○ ξ) (vec2list (scbwv_varstypes scbx))) as z. unfold scbwv_varstypes in z. rewrite vec2list_map_list2vec in z. rewrite fst_zip in z. rewrite <- z. clear z. + unfold sac_gamma in *. + simpl in *. + Unset Printing Implicit. + idtac. + apply nd_rule. + apply RArrange. + set (scbwv_exprvars_distinct scbx) as q'. + rewrite <- leaves_unleaves in q'. + apply (AComp (@factorContextRightAndWeaken'' _ (weakCE' Δ) _ _ (expr2antecedent ex) q')). + clear q'. - replace (stripOutVars (vec2list (scbwv_exprvars scbx))) with - (stripOutVars (leaves (unleaves (vec2list (scbwv_exprvars scbx))))). - apply q. - apply (sac_delta sac Γ atypes (weakCK'' Δ)). + set (scbwv_coherent scbx l ξ) as H. rewrite leaves_unleaves. - apply (scbwv_exprvars_distinct scbx). - rewrite leaves_unleaves. - reflexivity. + unfold scbwv_varstypes. + apply ALeft. + rewrite <- mapleaves'. + rewrite <- mapleaves'. + rewrite mapleaves'. + rewrite vec2list_map_list2vec. + rewrite <- H. + clear H. + rewrite <- mapleaves'. + rewrite vec2list_map_list2vec. + unfold scbwv_xi. + unfold scbwv_varstypes. + apply AId. destruct case_nil. apply nd_id0.