X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FProgrammingLanguage.v;h=e29903a4524935a6f1d2f88460878e90a4f687a0;hp=657a69f216befb695d3518a36efa0138cde48468;hb=4e5aa4bcc6024aa7add9c1bf1c2ad9fd2a6a3685;hpb=61fb093700aab006b77998d1cbd30235e144ca1f diff --git a/src/ProgrammingLanguage.v b/src/ProgrammingLanguage.v index 657a69f..e29903a 100644 --- a/src/ProgrammingLanguage.v +++ b/src/ProgrammingLanguage.v @@ -1,7 +1,7 @@ (*********************************************************************************************************************************) (* ProgrammingLanguage *) (* *) -(* Basic assumptions about programming languages . *) +(* Basic assumptions about programming languages. *) (* *) (*********************************************************************************************************************************) @@ -18,46 +18,27 @@ Require Import Enrichment_ch2_8. Require Import Subcategories_ch7_1. Require Import NaturalTransformations_ch7_4. Require Import NaturalIsomorphisms_ch7_5. +Require Import BinoidalCategories. +Require Import PreMonoidalCategories. Require Import MonoidalCategories_ch7_8. Require Import Coherence_ch7_8. Require Import Enrichment_ch2_8. Require Import RepresentableStructure_ch7_2. Require Import FunctorCategories_ch7_7. +Require Import Enrichments. Require Import NaturalDeduction. Require Import NaturalDeductionCategory. -Require Import FreydCategories. - -Require Import Reification. -Require Import GeneralizedArrow. -Require Import GeneralizedArrowFromReification. -Require Import ReificationFromGeneralizedArrow. - -(* - * Everything in the rest of this section is just groundwork meant to - * build up to the definition of the ProgrammingLanguage class, which - * appears at the end of the section. References to "the instance" - * mean instances of that class. Think of this section as being one - * big Class { ... } definition, except that we declare most of the - * stuff outside the curly brackets in order to take advantage of - * Coq's section mechanism. - *) 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. - (* Because of term irrelevance we need only store the *erased* (def - * 4.4) trees; for this reason there is no Coq type directly - * corresponding to productions $e$ and $x$ of 4.1.1, and TreeOT can - * be used for productions $\Gamma$ and $\Sigma$ *) - (* to do: sequent calculus equals natural deduction over sequents, theorem equals sequent with null antecedent, *) - - Context {Rule : Tree ??Judg -> Tree ??Judg -> Type}. + Context {Rule : Tree ??PLJudg -> Tree ??PLJudg -> Type}. Notation "H /⋯⋯/ C" := (ND Rule H C) : pl_scope. @@ -65,30 +46,12 @@ Section Programming_Language. Open Scope nd_scope. Open Scope pl_scope. - (* - * - * Note that from this abstract interface, the terms (expressions) - * in the proof are not accessible at all; they don't need to be -- - * so long as we have access to the equivalence relation upon - * proof-conclusions. Moreover, hiding the expressions actually - * makes the encoding in CiC work out easier for two reasons: - * - * 1. Because the denotation function is provided a proof rather - * than a term, it is a total function (the denotation function is - * often undefined for ill-typed terms). - * - * 2. We can define arr_composition of proofs without having to know how - * to compose expressions. The latter task is left up to the client - * function which extracts an expression from a completed proof. - * - * This also means that we don't need an explicit proof obligation for 4.1.2. - *) 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. @@ -103,146 +66,199 @@ Section Programming_Language. Definition identityProof t : [] ~~{JudgmentsL}~~> [t |= t]. unfold hom; simpl. - apply nd_seq_reflexive. + apply snd_initial. Defined. Definition cutProof a b c : [a |= b],,[b |= c] ~~{JudgmentsL}~~> [a |= c]. unfold hom; simpl. - apply pl_subst. + apply snd_cut. Defined. + Existing Instance pl_eqv. + Definition TypesL : ECategory JudgmentsL (Tree ??T) (fun x y => [x|=y]). refine {| eid := identityProof ; ecomp := cutProof |}; intros. - apply MonoidalCat_all_central. - apply MonoidalCat_all_central. - 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. + apply (mon_commutative(MonoidalCat:=JudgmentsL)). + apply (mon_commutative(MonoidalCat:=JudgmentsL)). + unfold identityProof; unfold cutProof; simpl; eapply cndr_inert. apply pl_eqv. auto. auto. + unfold identityProof; unfold cutProof; simpl; eapply cndr_inert. apply pl_eqv. auto. auto. + unfold identityProof; unfold cutProof; simpl; eapply cndr_inert. apply pl_eqv. auto. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. Defined. - Definition Types_first c : EFunctor TypesL TypesL (fun x => x,,c ). - refine {| efunc := fun x y => (nd_rule (@se_expand_right _ _ _ _ _ _ _ (@pl_sequent_join PL) c x y)) |}. - intros; apply MonoidalCat_all_central. + Instance Types_first c : EFunctor TypesL TypesL (fun x => x,,c ) := + { efunc := fun x y => cnd_expand_right(ContextND:=pl_cnd) x y c }. + intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)). intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof. - apply se_reflexive_right. + apply (cndr_inert pl_cnd); auto. 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#]). + rewrite <- (@ndr_prod_preserves_comp _ _ pl_eqv _ _ (cnd_expand_right _ _ c) _ _ (nd_id1 (b|=c0)) + _ (nd_id1 (a,,c |= b,,c)) _ (cnd_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. + simpl; eapply cndr_inert. apply pl_eqv. auto. auto. Defined. - Definition Types_second c : EFunctor TypesL TypesL (fun x => c,,x). - eapply Build_EFunctor. - instantiate (1:=(fun x y => (nd_rule (@se_expand_left _ _ _ _ _ _ _ (@pl_sequent_join PL) c x y)))). - intros; apply MonoidalCat_all_central. + Instance Types_second c : EFunctor TypesL TypesL (fun x => c,,x) := + { efunc := fun x y => ((@cnd_expand_left _ _ _ _ _ _ x y c)) }. + intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)). intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof. - apply se_reflexive_left. + eapply cndr_inert; auto. apply pl_eqv. 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#]). + rewrite <- (@ndr_prod_preserves_comp _ _ pl_eqv _ _ (cnd_expand_left _ _ c) _ _ (nd_id1 (b|=c0)) + _ (nd_id1 (c,,a |= c,,b)) _ (cnd_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. + simpl; eapply cndr_inert. apply pl_eqv. auto. auto. Defined. - Definition Types_binoidal : BinoidalCat TypesL (@T_Branch _). + Definition Types_binoidal : EBinoidalCat TypesL. refine - {| bin_first := Types_first - ; bin_second := Types_second + {| ebc_first := Types_first + ; ebc_second := Types_second |}. Defined. - Definition Types_PreMonoidal : PreMonoidalCat Types_binoidal []. + Instance Types_assoc_iso a b c : Isomorphic(C:=TypesL) ((a,,b),,c) (a,,(b,,c)) := + { iso_forward := snd_initial _ ;; cnd_ant_cossa _ a b c + ; iso_backward := snd_initial _ ;; cnd_ant_assoc _ a b c + }. + simpl; eapply cndr_inert. unfold identityProof; apply pl_eqv. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. + auto. + simpl; eapply cndr_inert. unfold identityProof; apply pl_eqv. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. + auto. + Defined. + + Instance Types_cancelr_iso a : Isomorphic(C:=TypesL) (a,,[]) a := + { iso_forward := snd_initial _ ;; cnd_ant_rlecnac _ a + ; iso_backward := snd_initial _ ;; cnd_ant_cancelr _ a + }. + unfold eqv; unfold comp; simpl. + eapply cndr_inert. apply pl_eqv. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. + auto. + unfold eqv; unfold comp; simpl. + eapply cndr_inert. apply pl_eqv. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. + auto. + Defined. + + Instance Types_cancell_iso a : Isomorphic(C:=TypesL) ([],,a) a := + { iso_forward := snd_initial _ ;; cnd_ant_llecnac _ a + ; iso_backward := snd_initial _ ;; cnd_ant_cancell _ a + }. + unfold eqv; unfold comp; simpl. + eapply cndr_inert. apply pl_eqv. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. + auto. + unfold eqv; unfold comp; simpl. + eapply cndr_inert. apply pl_eqv. auto. + apply ndpc_comp; auto. + apply ndpc_comp; auto. + auto. + 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. - Definition TypesEnrichedInJudgments : Enrichment. - refine {| enr_c := TypesL |}. + Instance Types_cancelr : Types_first [] <~~~> functor_id _ := + { ni_iso := Types_cancelr_iso }. + intros; simpl. + admit. Defined. - Structure HasProductTypes := - { - }. + Instance Types_cancell : Types_second [] <~~~> functor_id _ := + { ni_iso := Types_cancell_iso }. + admit. + Defined. - (* need to prove that if we have cartesian tuples we have cartesian contexts *) - Definition LanguagesWithProductsAreSMME : HasProductTypes -> SurjectiveMonicMonoidalEnrichment TypesEnrichedInJudgments. + 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. - End LanguageCategory. + 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. - Structure ProgrammingLanguageSMME := - { plsmme_pl : ProgrammingLanguage - ; plsmme_smme : SurjectiveMonicMonoidalEnrichment (TypesEnrichedInJudgments plsmme_pl) - }. - Coercion plsmme_pl : ProgrammingLanguageSMME >-> ProgrammingLanguage. - Coercion plsmme_smme : ProgrammingLanguageSMME >-> SurjectiveMonicMonoidalEnrichment. - - Section ArrowInLanguage. - Context (Host:ProgrammingLanguageSMME). - Context `(CC:CartesianCat (me_mon Host)). - Context `(K:@ECategory _ _ _ _ _ _ (@car_mn _ _ _ _ _ _ _ CC) C Kehom). - Context `(pmc:PreMonoidalCat K bobj mobj (@one _ _ _ (cartesian_terminal C))). - (* FIXME *) - (* - Definition ArrowInProgrammingLanguage := - @FreydCategory _ _ _ _ _ _ (@car_mn _ _ _ _ _ _ _ CC) _ _ _ _ pmc. - *) - End ArrowInLanguage. - - Section GArrowInLanguage. - Context (Guest:ProgrammingLanguageSMME). - Context (Host :ProgrammingLanguageSMME). - Definition GeneralizedArrowInLanguage := GeneralizedArrow Guest Host. - - (* FIXME - Definition ArrowsAreGeneralizedArrows : ArrowInProgrammingLanguage -> GeneralizedArrowInLanguage. - *) - Definition TwoLevelLanguage := Reification Guest Host (me_i Host). - - Context (GuestHost:TwoLevelLanguage). - - Definition FlatObject (x:TypesL Host) := - forall y1 y2, not ((reification_r_obj GuestHost y1 y2)=x). - - Definition FlatSubCategory := FullSubcategory (TypesL Host) FlatObject. - - Section Flattening. - - Context (F:Retraction (TypesL Host) FlatSubCategory). - Definition FlatteningOfReification := garrow_from_reification Guest Host GuestHost >>>> F. - Lemma FlatteningIsNotDestructive : - FlatteningOfReification >>>> retraction_retraction F >>>> RepresentableFunctor _ (me_i Host) ~~~~ GuestHost. - admit. - Qed. - - End Flattening. - - End GArrowInLanguage. - - Inductive NLevelLanguage : nat -> ProgrammingLanguageSMME -> Type := - | NLevelLanguage_zero : forall lang, NLevelLanguage O lang - | NLevelLanguage_succ : forall (L1 L2:ProgrammingLanguageSMME) n, - TwoLevelLanguage L1 L2 -> NLevelLanguage n L1 -> NLevelLanguage (S n) L2. - - Definition OmegaLevelLanguage : Type := - { f : nat -> ProgrammingLanguageSMME - & forall n, TwoLevelLanguage (f n) (f (S n)) }. - - Close Scope temporary_scope3. - Close Scope pl_scope. - Close Scope nd_scope. - Close Scope pf_scope. + Instance TypesL_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 + }. + apply Build_Pentagon. + intros; simpl. + eapply cndr_inert. apply pl_eqv. + apply ndpc_comp. + apply ndpc_comp. + auto. + apply ndpc_comp. + apply ndpc_prod. + apply ndpc_comp. + apply ndpc_comp. + auto. + apply ndpc_comp. + auto. + auto. + auto. + auto. + auto. + auto. + apply ndpc_comp. + apply ndpc_comp. + auto. + apply ndpc_comp. + auto. + auto. + auto. + apply Build_Triangle; intros; simpl. + eapply cndr_inert. apply pl_eqv. + auto. + apply ndpc_comp. + apply ndpc_comp. + auto. + apply ndpc_comp. + auto. + auto. + auto. + eapply cndr_inert. apply pl_eqv. auto. + auto. + intros; simpl; reflexivity. + intros; simpl; reflexivity. + admit. (* assoc central *) + admit. (* cancelr central *) + admit. (* cancell central *) + Defined. -End Programming_Language. + Definition TypesEnrichedInJudgments : SurjectiveEnrichment. + refine + {| senr_c_pm := TypesL_PreMonoidal + ; senr_v := JudgmentsL + ; senr_v_bin := Judgments_Category_binoidal _ + ; senr_v_pmon := Judgments_Category_premonoidal _ + ; senr_v_mon := Judgments_Category_monoidal _ + ; senr_c_bin := Types_binoidal + ; senr_c := TypesL + |}. + Defined. + End LanguageCategory. + +End Programming_Language. Implicit Arguments ND [ Judgment ].