X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FProgrammingLanguage.v;h=83b435ab06ac2b44d545a7f2a5bb6c769d0a2078;hp=7831fadaa911815baf3b0471a1db4fc6b98e20dc;hb=423b0bd3972c5bcbbd757cb715e13b5b9104a9a6;hpb=c3b1fb9622a65ad01e54b6e35785cee672d25bdc diff --git a/src/ProgrammingLanguage.v b/src/ProgrammingLanguage.v index 7831fad..83b435a 100644 --- a/src/ProgrammingLanguage.v +++ b/src/ProgrammingLanguage.v @@ -27,17 +27,16 @@ Require Import RepresentableStructure_ch7_2. Require Import FunctorCategories_ch7_7. Require Import NaturalDeduction. -Require Import NaturalDeductionCategory. Section Programming_Language. Context {T : Type}. (* types of the language *) - Context (Judg : Type). - Context (sequent : Tree ??T -> Tree ??T -> Judg). + Definition PLJudg := (Tree ??T) * (Tree ??T). + Definition sequent := @pair (Tree ??T) (Tree ??T). Notation "cs |= ss" := (sequent cs ss) : pl_scope. - Context {Rule : Tree ??Judg -> Tree ??Judg -> Type}. + Context {Rule : Tree ??PLJudg -> Tree ??PLJudg -> Type}. Notation "H /⋯⋯/ C" := (ND Rule H C) : pl_scope. @@ -46,180 +45,15 @@ Section Programming_Language. Open Scope pl_scope. Class ProgrammingLanguage := - { pl_eqv : @ND_Relation Judg Rule where "pf1 === pf2" := (@ndr_eqv _ _ pl_eqv _ _ pf1 pf2) - ; pl_tsr :> @TreeStructuralRules Judg Rule T sequent - ; pl_sc :> @SequentCalculus Judg Rule _ sequent - ; pl_subst :> @CutRule Judg Rule _ sequent pl_eqv pl_sc - ; pl_sequent_join :> @SequentExpansion Judg Rule T sequent pl_eqv pl_sc pl_subst + { pl_eqv0 :> @ND_Relation PLJudg Rule + ; pl_snd :> @SequentND PLJudg Rule _ sequent + ; pl_cnd :> @ContextND PLJudg Rule T sequent pl_snd + ; pl_eqv1 :> @SequentND_Relation PLJudg Rule _ sequent pl_snd pl_eqv0 + ; pl_eqv :> @ContextND_Relation PLJudg Rule _ sequent pl_snd pl_cnd pl_eqv0 pl_eqv1 }. Notation "pf1 === pf2" := (@ndr_eqv _ _ pl_eqv _ _ pf1 pf2) : temporary_scope3. - - Section LanguageCategory. - - Context (PL:ProgrammingLanguage). - - (* category of judgments in a fixed type/coercion context *) - Definition Judgments_cartesian := @Judgments_Category_CartesianCat _ Rule pl_eqv. - - Definition JudgmentsL := Judgments_cartesian. - - Definition identityProof t : [] ~~{JudgmentsL}~~> [t |= t]. - unfold hom; simpl. - apply nd_seq_reflexive. - Defined. - - Definition cutProof a b c : [a |= b],,[b |= c] ~~{JudgmentsL}~~> [a |= c]. - unfold hom; simpl. - apply pl_subst. - Defined. - - Existing Instance pl_eqv. - Definition TypesL : ECategory JudgmentsL (Tree ??T) (fun x y => [x|=y]). - refine - {| eid := identityProof - ; ecomp := cutProof - |}; intros. - apply (mon_commutative(MonoidalCat:=JudgmentsL)). - apply (mon_commutative(MonoidalCat:=JudgmentsL)). - unfold identityProof; unfold cutProof; simpl. - apply nd_cut_left_identity. - unfold identityProof; unfold cutProof; simpl. - apply nd_cut_right_identity. - unfold identityProof; unfold cutProof; simpl. - symmetry. - apply nd_cut_associativity. - Defined. - - Definition Types_first c : EFunctor TypesL TypesL (fun x => x,,c ). - refine {| efunc := fun x y => (@se_expand_right _ _ _ _ _ _ _ (@pl_sequent_join PL) c x y) |}. - intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)). - intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof. - apply se_reflexive_right. - intros. unfold ehom. unfold comp. simpl. unfold cutProof. - rewrite <- (@ndr_prod_preserves_comp _ _ pl_eqv _ _ (se_expand_right _ c) _ _ (nd_id1 (b|=c0)) - _ (nd_id1 (a,,c |= b,,c)) _ (se_expand_right _ c)). - setoid_rewrite (@ndr_comp_right_identity _ _ pl_eqv _ [a,, c |= b,, c]). - setoid_rewrite (@ndr_comp_left_identity _ _ pl_eqv [b |= c0]). - apply se_cut_right. - Defined. - - Definition Types_second c : EFunctor TypesL TypesL (fun x => c,,x). - eapply Build_EFunctor. - instantiate (1:=(fun x y => ((@se_expand_left _ _ _ _ _ _ _ (@pl_sequent_join PL) c x y)))). - intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)). - intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof. - apply se_reflexive_left. - intros. unfold ehom. unfold comp. simpl. unfold cutProof. - rewrite <- (@ndr_prod_preserves_comp _ _ pl_eqv _ _ (se_expand_left _ c) _ _ (nd_id1 (b|=c0)) - _ (nd_id1 (c,,a |= c,,b)) _ (se_expand_left _ c)). - setoid_rewrite (@ndr_comp_right_identity _ _ pl_eqv _ [c,,a |= c,,b]). - setoid_rewrite (@ndr_comp_left_identity _ _ pl_eqv [b |= c0]). - apply se_cut_left. - Defined. - - Definition Types_binoidal : BinoidalCat TypesL (@T_Branch _). - refine - {| bin_first := Types_first - ; bin_second := Types_second - |}. - Defined. - - Instance Types_assoc_iso a b c : Isomorphic(C:=TypesL) ((a,,b),,c) (a,,(b,,c)) := - { iso_forward := nd_seq_reflexive _ ;; tsr_ant_cossa _ a b c - ; iso_backward := nd_seq_reflexive _ ;; tsr_ant_assoc _ a b c - }. - admit. - admit. - Defined. - - Instance Types_assoc a b : Types_second a >>>> Types_first b <~~~> Types_first b >>>> Types_second a := - { ni_iso := fun c => Types_assoc_iso a c b }. - admit. - Defined. - - Instance Types_cancelr_iso a : Isomorphic(C:=TypesL) (a,,[]) a := - { iso_forward := nd_seq_reflexive _ ;; tsr_ant_rlecnac _ a - ; iso_backward := nd_seq_reflexive _ ;; tsr_ant_cancelr _ a - }. - admit. - admit. - Defined. - - Instance Types_cancelr : Types_first [] <~~~> functor_id _ := - { ni_iso := Types_cancelr_iso }. - admit. - Defined. - - Instance Types_cancell_iso a : Isomorphic(C:=TypesL) ([],,a) a := - { iso_forward := nd_seq_reflexive _ ;; tsr_ant_llecnac _ a - ; iso_backward := nd_seq_reflexive _ ;; tsr_ant_cancell _ a - }. - admit. - admit. - Defined. - - Instance Types_cancell : Types_second [] <~~~> functor_id _ := - { ni_iso := Types_cancell_iso }. - admit. - Defined. - - Instance Types_assoc_ll a b : Types_second (a,,b) <~~~> Types_second b >>>> Types_second a := - { ni_iso := fun c => Types_assoc_iso a b c }. - admit. - Defined. - - Instance Types_assoc_rr a b : Types_first (a,,b) <~~~> Types_first a >>>> Types_first b := - { ni_iso := fun c => iso_inv _ _ (Types_assoc_iso c a b) }. - admit. - Defined. - - Instance Types_PreMonoidal : PreMonoidalCat Types_binoidal [] := - { pmon_assoc := Types_assoc - ; pmon_cancell := Types_cancell - ; pmon_cancelr := Types_cancelr - ; pmon_assoc_rr := Types_assoc_rr - ; pmon_assoc_ll := Types_assoc_ll - }. - admit. (* pentagon law *) - admit. (* triangle law *) - intros; simpl; reflexivity. - intros; simpl; reflexivity. - admit. (* assoc central *) - admit. (* cancelr central *) - admit. (* cancell central *) - Defined. - - (* - Definition TypesEnrichedInJudgments : Enrichment. - refine {| enr_c := TypesL |}. - Defined. - - Structure HasProductTypes := - { - }. - - Lemma CartesianEnrMonoidal (e:Enrichment) `(C:CartesianCat(Ob:= _)(Hom:= _)(C:=Underlying (enr_c e))) : MonoidalEnrichment e. - admit. - Defined. - - (* need to prove that if we have cartesian tuples we have cartesian contexts *) - Definition LanguagesWithProductsAreSMME : HasProductTypes -> SurjectiveMonicMonoidalEnrichment TypesEnrichedInJudgments. - admit. - Defined. - *) - End LanguageCategory. + Coercion pl_eqv : ProgrammingLanguage >-> ContextND_Relation. + Coercion pl_cnd : ProgrammingLanguage >-> ContextND. End Programming_Language. -(* -Structure ProgrammingLanguageSMME := -{ plsmme_t : Type -; plsmme_judg : Type -; plsmme_sequent : Tree ??plsmme_t -> Tree ??plsmme_t -> plsmme_judg -; plsmme_rule : Tree ??plsmme_judg -> Tree ??plsmme_judg -> Type -; plsmme_pl : @ProgrammingLanguage plsmme_t plsmme_judg plsmme_sequent plsmme_rule -; plsmme_smme : SurjectiveMonicMonoidalEnrichment (TypesEnrichedInJudgments _ _ plsmme_pl) -}. -Coercion plsmme_pl : ProgrammingLanguageSMME >-> ProgrammingLanguage. -Coercion plsmme_smme : ProgrammingLanguageSMME >-> SurjectiveMonicMonoidalEnrichment. -*) -Implicit Arguments ND [ Judgment ]. +