Library Progress_proofs
Require
Import
List.
Require
Import
Util.
Require
Syntax
.
Module
SetProgram(My_Program: Syntax.Program).
Import
My_Program.
Require
TransitivityElimination_proofs.
Module
My_TransitivityElimination_proofs := TransitivityElimination_proofs.SetProgram(My_Program).
Import
My_TransitivityElimination_proofs.
Import
My_Normalization_proofs.
Import
My_Subtyping_proofs.
Import
My_Typing_proofs.
Import
My_Confluence.
Import
My_Auxiliary_proofs.
Import
My_Substitutions_proofs.
Import
My_Binders_proofs.
Import
My_Semantics.
Import
My_WellFormedness.
Import
My_Typing.
Import
My_Auxiliary.
Import
My_Substitutions.
Import
My_Binders.
Lemma
evaluated_term_dec:
forall (t: Term),
(evaluated_term t) \/ ~(evaluated_term t)
.
Lemma
Progress_aux_1:
forall (f: FieldSym) (fts: Term),
(In f (fields_of_fieldTermList fts)) ->
exists t: Term,
(fieldTerm_lookup fts f t)
.
Lemma
evaluated_term_new_inversion:
forall (D: TypeEnv) (G: TermEnv) (t: Term) (T: TypeS),
(Term_WT D G t (TermI T)) ->
(evaluated_term t) ->
exists N: TypeS, exists fts: Term,
t = (Term_new N fts) /\
(Term_WT D G fts (FieldTermListI N))
.
Lemma
KindS_at_is_a_partial_function:
forall (ks k k': KindS) (i: nat),
(KindS_at ks i k) ->
(KindS_at ks i k') ->
(k = k')
.
Inductive
SameCategory: TypeC -> TypeC -> Prop :=
| SameCategory_TypeConstr:
forall (k k': KindS),
(SameCategory (TypeConstr k) (TypeConstr k'))
| SameCategory_Type_:
(SameCategory Type_ Type_)
| SameCategory_TypeTuple:
forall (k k': KindS),
(SameCategory (TypeTuple k) (TypeTuple k'))
| SameCategory_TypeParam:
(SameCategory TypeParam TypeParam)
| SameCategory_TypeSection:
(SameCategory TypeSection TypeSection)
| SameCategory_TypeOption:
(SameCategory TypeOption TypeOption)
| SameCategory_TypeParamList:
(SameCategory TypeParamList TypeParamList)
| SameCategory_TypeConstrList:
forall (k k': KindS),
(SameCategory (TypeConstrList k) (TypeConstrList k'))
.
Lemma
Uniqueness_of_well_kindedness:
forall (D: TypeEnv) (t: TypeS) (c1 c2: TypeC),
(Type_WK D t c1) ->
(Type_WK D t c2) ->
(SameCategory c1 c2) ->
(c1 = c2)
.
Lemma
Satisfaction_implies_same_kind:
forall (D: TypeEnv) (R Pi: TypeS) (k: KindS),
(Type_WK D R (TypeTuple k)) ->
(SectionSat D R Pi) ->
(kind_of Pi k)
.
Lemma
Substitution_in_field_type:
forall (D: TypeEnv) (f: FieldSym) (R: TypeS),
let C := (fieldOwner f) in
let T := (fieldType f) in
(Type_WK D (Class_type C R) Type_) ->
(Type_WK D (typeSubst_in_type T 0 R) Type_)
.
Proof
.
intros.
inversion H. clear D0 K R0 H2 H0 H1 H.
inversion H3. clear D0 C0 k0 H0 H H1 H3.
destruct (field_declarations_are_well_formed f).
eapply Substitution_preserves_well_kindedness_snoc. apply H4.
apply H2.
replace (snoc D (classTypeParams C)) with (D ++ (Pi :: nil)).
apply Weakening_of_well_kindedness_by_the_left. trivial.
trivial.
Qed
.
Lemma
Substitution_in_method_type_params:
forall (D: TypeEnv) (m: MethodSym) (R: TypeS),
let C := (methodOwner m) in
let Pim := (methodTypeParams m) in
(Type_WK D (Class_type C R) Type_) ->
(Type_WK D (typeSubst_in_type Pim 0 R) TypeSection)
.
Proof
.
intros.
inversion H. clear D0 K R0 H2 H0 H1 H.
inversion H3. clear D0 C0 k0 H0 H H1 H3.
destruct (method_declarations_are_well_formed m).
eapply Substitution_preserves_well_kindedness_snoc. eauto. eauto.
change (snoc D (classTypeParams C)) with (D ++ (PiC :: nil)).
apply Weakening_of_well_kindedness_by_the_left; trivial.
Qed
.
Lemma
Substitution_in_method_type:
forall (D: TypeEnv) (m: MethodSym) (R R': TypeS),
let C := (methodOwner m) in
let T := (methodType m) in
let Pi := (methodTypeParams m) in
(Type_WK D (Class_type C R) Type_) ->
(SectionSat D R' (typeSubst_in_type Pi 0 R)) ->
(Type_WK D (typeSubst_in_type (typeSubst_in_type T 1 R) 0 R') Type_)
.
Proof
.
intros.
inversion H. clear D0 K R0 H3 H1 H2 H.
inversion H4. clear D0 C0 k0 H1 H H2 H4.
inversion H0. clear H4 H0.
destruct (method_declarations_are_well_formed m). clear H4.
assert (kind_of (typeSubst_in_type Pim 0 R) (Kind_tuple ks)).
replace Pim with Pi. rewrite <- H. apply Kind_of_section. trivial. trivial.
eapply Substitution_preserves_well_kindedness_snoc. apply H2. apply H4.
replace (snoc D (typeSubst_in_type Pim 0 R)) with
(D ++ (typeSubst_in_typeEnv (Pim :: nil) 0 R)).
replace 1 with (length (Pim :: nil)).
eapply Type_substitution_preserves_well_kindedness
with (D:= D ++ PiC :: Pim :: nil).
apply H5. apply H3.
apply Weakening_of_well_kindedness_by_the_left.
apply H6. trivial. trivial.
simpl. trivial.
Qed
.
Lemma
Substitution_in_method_param_types:
forall (D: TypeEnv) (m: MethodSym) (R R': TypeS),
let C := (methodOwner m) in
let Ts := (methodParams m) in
let Pi := (methodTypeParams m) in
(Type_WK D (Class_type C R) Type_) ->
(SectionSat D R' (typeSubst_in_type Pi 0 R)) ->
(List_forall
(fun T => (Type_WK D (typeSubst_in_type (typeSubst_in_type T 1 R) 0 R') Type_))
Ts)
.
Proof
.
intros.
inversion H. clear D0 K R0 H3 H1 H2 H.
inversion H4. clear D0 C0 k0 H1 H H2 H4.
inversion H0.
assert (kind_of (typeSubst_in_type Pi 0 R) (Kind_tuple ks)).
eapply Satisfaction_implies_same_kind; eauto.
destruct (method_declarations_are_well_formed m).
eapply List_forall_implies_forall with
(P1:= (fun U : TypeS => Type_WK D0 U Type_)).
intros. rename x into U.
eapply Substitution_preserves_well_kindedness_snoc. apply H2. apply H6.
replace (snoc D (typeSubst_in_type Pi 0 R)) with
(D ++ (typeSubst_in_typeEnv (Pi :: nil) 0 R)).
replace 1 with (length (Pi :: nil)).
eapply Type_substitution_preserves_well_kindedness
with (D:= D ++ PiC :: Pi :: nil).
apply H5. apply H3.
apply Weakening_of_well_kindedness_by_the_left.
apply H10. trivial. trivial.
simpl. trivial. trivial.
Qed
.
Lemma
Substitution_in_method_param_types_2:
forall (D: TypeEnv) (m: MethodSym) (R: TypeS),
let C := (methodOwner m) in
let Ts := (methodParams m) in
let Pim := (methodTypeParams m) in
(Type_WK D (Class_type C R) Type_) ->
(List_forall
(fun T => (Type_WK (D ++ (typeSubst_in_type Pim 0 R) :: nil)
(typeSubst_in_type T 1 R) Type_))
Ts)
.
Proof
.
intros.
inversion H. clear D0 K R0 H2 H0 H1 H.
inversion H3. clear D0 C0 k0 H0 H H1 H3.
destruct (method_declarations_are_well_formed m).
eapply List_forall_implies_forall with
(P1:= (fun U : TypeS => Type_WK D0 U Type_)).
intros. rename x into U. unfold D0 in H3.
assert (Type_WK (D ++ (PiC :: (Pim :: nil))) U Type_).
apply Weakening_of_well_kindedness_by_the_left. trivial.
clear H3.
change (typeSubst_in_type Pim 0 R :: nil) with
(typeSubst_in_typeEnv (Pim :: nil) 0 R).
change 1 with (length (Pim :: nil)).
eapply Type_substitution_preserves_well_kindedness; eauto.
trivial.
Qed
.
Lemma
Substitution_in_method_type_2:
forall (D: TypeEnv) (m: MethodSym) (R: TypeS),
let C := (methodOwner m) in
let T := (methodType m) in
let Pim := (methodTypeParams m) in
(Type_WK D (Class_type C R) Type_) ->
(Type_WK (D ++ (typeSubst_in_type Pim 0 R) :: nil)
(typeSubst_in_type T 1 R) Type_)
.
Proof
.
intros.
inversion H. clear D0 K R0 H2 H0 H1 H.
inversion H3. clear D0 C0 k0 H0 H H1 H3.
destruct (method_declarations_are_well_formed m).
assert (Type_WK (D ++ (PiC :: (Pim :: nil))) T Type_).
apply Weakening_of_well_kindedness_by_the_left. trivial.
change (typeSubst_in_type Pim 0 R :: nil) with
(typeSubst_in_typeEnv (Pim :: nil) 0 R).
change 1 with (length (Pim :: nil)).
eapply Type_substitution_preserves_well_kindedness; eauto.
Qed
.
Definition
Info_WK (D: TypeEnv) (info: TermInfo): Prop :=
match info with
| (TermI T) => (Type_WK D T Type_)
| (TermListI Ts) => List_forall (fun T => (Type_WK D T Type_)) Ts
| (FieldTermListI N) => True
end.
Theorem
Typing_returns_well_kinded_types:
forall (D: TypeEnv) (G: TermEnv) (t: Term) (info: TermInfo),
(Term_WT D G t info) ->
(List_forall (fun T => (Type_WK D T Type_)) G) ->
(Info_WK D info)
.
Lemma
Typing_returns_well_kinded_types_cor1:
forall (D: TypeEnv) (G: TermEnv) (t: Term) (T: TypeS),
(Term_WT D G t (TermI T)) ->
(List_forall (fun T => (Type_WK D T Type_)) G) ->
(Type_WK D T Type_)
.
Theorem
Progress:
forall (t: Term) (info: TermInfo),
(Term_WT nil nil t info) ->
~(evaluated_term t) ->
exists u: Term,
(RedTerm t u)
.
End
SetProgram.
Index
This page has been generated by coqdoc