Library SubjectReduction_proofs
Require
Import
List.
Require
Import
Arith.
Require
Import
MatchCompNat.
Require
Import
Util.
Require
Syntax
.
Module
SetProgram(My_Program: Syntax.Program).
Import
My_Program.
Require
SubjectReduction_preliminaries.
Module
My_SubjectReduction_preliminaries := SubjectReduction_preliminaries.SetProgram(My_Program).
Import
My_SubjectReduction_preliminaries.
Import
My_TermSubstitutionInTyping.
Import
My_TypeSubstitutionInTyping.
Import
My_BoundCompliantSubstitutions_proofs.
Import
My_Progress_proofs.
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
Typed_term_list:
forall (D: TypeEnv) (G: TermEnv) (ts: Term) (Ts: (list TypeS)),
(Term_WT D G ts (TermListI Ts)) ->
(Forall2 (fun t T => (Term_WT D G t (TermI T))) (list_of_termList ts) Ts)
.
Inductive
FieldVal_WF (D: TypeEnv) (G: TermEnv) (N: TypeS)
(f: FieldSym) (t: Term): Prop :=
| FieldVal_WF_intro:
forall (S R: TypeS),
let C := (fieldOwner f) in
let T := (fieldType f) in
(Type_WK D N Type_) ->
(Subtyping D N (Class_type C R)) ->
(Type_WK D (Class_type C R) Type_) ->
(Term_WT D G t (TermI S)) ->
(Subtyping D S (typeSubst_in_type T 0 R)) ->
(FieldVal_WF D G N f t)
.
Lemma
Well_formed_field_value:
forall (D: TypeEnv) (G: TermEnv) (N: TypeS) (f: FieldSym) (t fts: Term),
(Term_WT D G fts (FieldTermListI N)) ->
(fieldTerm_lookup fts f t) ->
(FieldVal_WF D G N f t)
.
Lemma
Term_reduction_preserves_fields_of_fieldTermList:
forall (D: TypeEnv) (G: TermEnv) (N: TypeS) (fts fts': Term),
(Term_WT D G fts (FieldTermListI N)) ->
(RedTerm fts fts') ->
(fields_of_fieldTermList fts') =
(fields_of_fieldTermList fts)
.
Lemma
Class_subtyping_implies_subtyping:
forall (D: TypeEnv) (C C': ClassSym) (R R': TypeS),
(Class_subtyping (Class_type C R) (Class_type C' R')) ->
(Subtyping D (Class_type C R) (Class_type C' R'))
.
Lemma
Self_type_is_well_kinded:
forall (C: ClassSym),
let Pi := (classTypeParams C) in
(Type_WK (Pi :: nil) (Class_type C (Type_var 0)) Type_)
.
Proof
.
intros. unfold Class_type.
destruct (Class_kind_existence C). rename x into k.
eapply WK_apply. apply WK_class. apply H.
eapply WK_var.
apply decl_of_typeVar_intro.
replace (Pi :: nil) with (snoc nil Pi).
apply TypeSection_at_tail. trivial.
replace (lift_in_type Pi 1) with Pi. trivial.
symmetry. unfold lift_in_type.
destruct (classes_are_well_kinded C).
eapply Lifting_in_closed_type_cor1; eauto.
Qed
.
Lemma
Type_env_reduction_preserves_typing:
forall (D D': TypeEnv) (G: TermEnv) (t: Term) (info: TermInfo),
(RedTypeEnvStar D D') ->
(Term_WT D G t info) ->
(Term_WT D' G t info)
.
Lemma
RTS_class_type:
forall (C: ClassSym) (R R': TypeS),
(RedTypeStar R R') ->
(RedTypeStar (Class_type C R) (Class_type C R'))
.
Lemma
RTE_head_star:
forall (D: TypeEnv) (T T': TypeS),
(RedTypeStar T T') ->
(RedTypeEnvStar (T :: D) (T' :: D))
.
Lemma
RTE_tail_star:
forall (D D': TypeEnv) (T: TypeS),
(RedTypeEnvStar D D') ->
(RedTypeEnvStar (T :: D) (T :: D'))
.
Lemma
Subject_reduction_case_T_call_redex_aux1:
forall (m: MethodSym) (Ci: ClassSym) (R Ri Ri0 R' Rm S: TypeS),
let C := (methodOwner m) in
let Pim := (methodTypeParams m) in
let T := (methodType m) in
let PiCi := (classTypeParams Ci) in
(Type_WK nil (Class_type C R) Type_) ->
(Type_WK (PiCi :: nil) (Class_type C Ri0) Type_) ->
(Type_WK nil (Class_type C (typeSubst_in_type Ri0 0 Ri)) Type_) ->
(Conform nil Ci Ri) ->
(SectionSat nil Rm (typeSubst_in_type Pim 0 R)) ->
(SectionSat nil Rm (typeSubst_in_type Pim 0 R')) ->
(RedTypeStar R R') ->
(RedTypeStar (typeSubst_in_type Ri0 0 Ri) R') ->
let D := (PiCi :: (typeSubst_in_type Pim 0 Ri0) :: nil) in
(Type_WK D S Type_) ->
(Subtyping D S (typeSubst_in_type T 1 Ri0)) ->
(Subtyping nil
(typeSubst_in_type (typeSubst_in_type S 1 Ri) 0 Rm)
(typeSubst_in_type (typeSubst_in_type T 1 R) 0 Rm))
.
Proof
.
intros until PiCi.
intros Hy1 Hy2 Hy3 Hy4 Hy5 H Hy6 Hy7.
intros D Hy8 Hy9.
inversion H.
eapply Type_substitution_preserves_subtyping_cor2. eauto. apply H.
change (typeSubst_in_type Pim 0 R' :: nil) with
(snoc nil (typeSubst_in_type Pim 0 R')). apply WK_typeEnv_snoc.
apply WK_typeEnv_nil.
unfold Pim. eapply Substitution_in_method_type_params.
eapply Well_kindedness_is_preserved_by_type_reduction_star.
apply Hy1.
apply RTS_class_type; trivial.
apply Well_kindedness_is_preserved_by_type_env_reduction_star with
(D:= (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri)) :: nil).
replace Pim with (typeSubst_in_type Pim 1 Ri).
Focus 2. unfold Pim. apply methodTypeParams_closed.
rewrite <- Permuting_type_substitutions_cor1.
inversion Hy4. inversion H4. inversion H9.
eapply Type_substitution_preserves_well_kindedness_cor1; eauto.
apply RTE_head_star. apply Reduction_in_substitution_star; trivial.
apply Well_kindedness_is_preserved_by_type_env_reduction_star with
(D:= (typeSubst_in_type Pim 0 R :: nil)).
change (typeSubst_in_type Pim 0 R :: nil) with
(nil ++ ((typeSubst_in_type Pim 0 R :: nil))).
unfold Pim. unfold T. apply Substitution_in_method_type_2. trivial.
apply RTE_head_star. apply Reduction_in_substitution_star; trivial.
apply S_beta_right_star with (U':= (typeSubst_in_type T 1 R')).
Focus 2.
apply Reduction_in_substitution_star; trivial.
apply Reduction_preserves_subtyping with
(D:= (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri)) :: nil)
(T:= (typeSubst_in_type S 1 Ri))
(U:= (typeSubst_in_type T 1 (typeSubst_in_type Ri0 0 Ri))).
Focus 2.
apply RTE_head_star.
apply Reduction_in_substitution_star. trivial.
Focus 2. apply RedType_refl.
Focus 2. apply Reduction_in_substitution_star. trivial.
replace Pim with (typeSubst_in_type Pim 1 Ri).
Focus 2.
unfold Pim. apply methodTypeParams_closed.
replace T with (typeSubst_in_type T 2 Ri).
Focus 2.
unfold T. apply methodType_closed.
rewrite <- Permuting_type_substitutions_cor1.
rewrite <- Permuting_type_substitutions_cor2.
inversion Hy4. inversion H4. clear D0 K R0 H8 H6 H7.
eapply Type_substitution_preserves_subtyping_cor1. eauto. eauto.
Focus 4. trivial.
fold PiCi.
replace (PiCi :: (typeSubst_in_type Pim 0 Ri0) :: nil) with
(snoc (snoc nil PiCi) (typeSubst_in_type Pim 0 Ri0)).
apply WK_typeEnv_snoc. apply WK_typeEnv_snoc. apply WK_typeEnv_nil.
destruct (classes_are_well_kinded Ci). trivial.
unfold Pim. eapply Substitution_in_method_type_params. eauto. trivial.
trivial.
fold PiCi.
change (PiCi :: typeSubst_in_type Pim 0 Ri0 :: nil) with
((PiCi :: nil) ++ typeSubst_in_type Pim 0 Ri0 :: nil).
unfold Pim. unfold T.
apply Substitution_in_method_type_2. trivial.
Qed
.
Lemma
Subject_reduction_case_T_call_redex_aux2:
forall (m: MethodSym) (Ci: ClassSym) (R Ri Ri0 R' Rm S: TypeS) (tb: Term),
let C := (methodOwner m) in
let Pim := (methodTypeParams m) in
let Ts := (methodParams m) in
let PiCi := (classTypeParams Ci) in
(Type_WK nil (Class_type C R) Type_) ->
(Type_WK (PiCi :: nil) (Class_type C Ri0) Type_) ->
(Type_WK nil (Class_type C (typeSubst_in_type Ri0 0 Ri)) Type_) ->
(Conform nil Ci Ri) ->
(SectionSat nil Rm (typeSubst_in_type Pim 0 R')) ->
(RedTypeStar R R') ->
(RedTypeStar (typeSubst_in_type Ri0 0 Ri) R') ->
let D := (PiCi :: (typeSubst_in_type Pim 0 Ri0) :: nil) in
let G := ((Class_type Ci (Type_var 1)) :: (typeSubst_in_termEnv Ts 1 Ri0)) in
(TermEnv_WK D G) ->
(Term_WT D G tb (TermI S)) ->
let G0 := ((Class_type Ci Ri) ::
(typeSubst_in_termEnv
(typeSubst_in_termEnv Ts 1 (typeSubst_in_type Ri0 0 Ri)) 0 Rm)) in
(Term_WT nil G0
(typeSubst_in_term (typeSubst_in_term tb 1 Ri) 0 Rm)
(TermI (typeSubst_in_type (typeSubst_in_type S 1 Ri) 0 Rm)))
.
Proof
.
intros until PiCi.
intros Hy1 Hy2 Hy3 Hy4 Hy5 Hy6 Hy7.
intros D G Hy42 Hy8 G0.
set (D':= (typeSubst_in_type (typeSubst_in_type Pim 0 Ri0) 0 Ri) :: nil).
set (G':= (typeSubst_in_type (Class_type Ci (Type_var 1)) 1 Ri) ::
(typeSubst_in_termEnv (typeSubst_in_termEnv Ts 1 Ri0) 1 Ri)).
assert (Term_WT D' G' (typeSubst_in_term tb 1 Ri) (TermI (typeSubst_in_type S 1 Ri))).
unfold D'. unfold G'.
inversion Hy4. inversion H. clear D0 K R0 H3 H1 H2.
eapply Type_substitution_preserves_typing_cor4; eauto.
fold PiCi.
replace (PiCi :: (typeSubst_in_type Pim 0 Ri0) :: nil) with
(snoc (snoc nil PiCi) (typeSubst_in_type Pim 0 Ri0)).
apply WK_typeEnv_snoc. apply WK_typeEnv_snoc. apply WK_typeEnv_nil.
destruct (classes_are_well_kinded Ci). trivial.
unfold Pim. eapply Substitution_in_method_type_params. apply Hy2. trivial.
clear D G Hy8 Hy42. rename D' into D. rename G' into G. rename H into Hy8.
assert ((typeSubst_in_type (Class_type Ci (Type_var 1)) 1 Ri) = (Class_type Ci Ri)).
simpl.
unfold typeSubst_in_typeVar. rewrite match_comp_nat3_eq.
replace (lift_in_type Ri 1) with Ri. trivial.
symmetry. unfold lift_in_type.
inversion Hy4. inversion H0.
eapply Lifting_in_closed_type_cor1; eauto. trivial.
rename H into Hy9.
unfold G in Hy8. rewrite Hy9 in Hy8. clear Hy9. unfold D in Hy8. clear D G.
rewrite Permuting_type_substitutions_cor1 in Hy8.
unfold Pim in Hy8. rewrite methodTypeParams_closed in Hy8. fold Pim in Hy8.
unfold G0.
replace (Class_type Ci Ri) with (typeSubst_in_type (Class_type Ci Ri) 0 Rm).
Focus 2.
inversion Hy4.
change 0 with ((length (nil (A:= TypeS))) + 0).
eapply Substitution_in_closed_type_aux; eauto.
change (typeSubst_in_type (Class_type Ci Ri) 0 Rm
:: typeSubst_in_termEnv
(typeSubst_in_termEnv Ts 1 (typeSubst_in_type Ri0 0 Ri)) 0 Rm) with
(typeSubst_in_termEnv ((Class_type Ci Ri) ::
(typeSubst_in_termEnv Ts 1 (typeSubst_in_type Ri0 0 Ri))) 0 Rm).
inversion Hy5.
eapply Type_substitution_preserves_typing_cor5; eauto.
change (typeSubst_in_type Pim 0 R' :: nil) with
(snoc nil (typeSubst_in_type Pim 0 R')).
constructor. constructor.
unfold Pim. eapply Substitution_in_method_type_params.
eapply Well_kindedness_is_preserved_by_type_reduction_star.
apply Hy1.
apply RTS_class_type; trivial.
constructor.
rewrite app_nil_end with (l:= (typeSubst_in_type Pim 0 R' :: nil)).
apply Weakening_of_well_kindedness_by_the_left. inversion Hy4. trivial.
unfold typeSubst_in_termEnv. apply List_forall_of_map.
apply List_forall_implies_forall with (P1:= fun U =>
(Type_WK ((typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri)) :: nil)
(typeSubst_in_type U 1 (typeSubst_in_type Ri0 0 Ri)) Type_)).
intros. rename x into U.
eapply Well_kindedness_is_preserved_by_type_env_reduction_star. eauto.
apply RTE_head_star. apply Reduction_in_substitution_star; trivial.
change (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri) :: nil) with
(nil ++ (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri) :: nil)).
unfold Pim. unfold Ts.
apply Substitution_in_method_param_types_2. trivial.
replace Ts with (typeSubst_in_termEnv Ts 2 Ri).
Focus 2.
unfold Ts. apply methodParams_closed.
rewrite <- Permuting_type_substitutions_in_lists_cor2.
change typeSubst_in_typeList with typeSubst_in_termEnv.
apply Type_env_reduction_preserves_typing with
(D:= (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri) :: nil)).
apply RTE_head_star. apply Reduction_in_substitution_star; trivial.
trivial.
Qed
.
Lemma
Subject_reduction_case_T_call_redex_aux3:
forall (m: MethodSym) (Ci: ClassSym) (R Ri Ri0 R' Rm S N: TypeS) (tb t ts: Term)
(Ss: (list TypeS)),
let C := (methodOwner m) in
let Pim := (methodTypeParams m) in
let Ts := (methodParams m) in
let PiCi := (classTypeParams Ci) in
(Type_WK nil (Class_type C R) Type_) ->
(Type_WK (PiCi :: nil) (Class_type C Ri0) Type_) ->
(Type_WK nil (Class_type C (typeSubst_in_type Ri0 0 Ri)) Type_) ->
(Conform nil Ci Ri) ->
(SectionSat nil Rm (typeSubst_in_type Pim 0 R')) ->
(RedTypeStar R R') ->
(RedTypeStar (typeSubst_in_type Ri0 0 Ri) R') ->
(Term_WT nil nil t (TermI N)) ->
(Subtyping nil N (Class_type Ci Ri)) ->
let Ts1 := (typeSubst_in_termEnv (typeSubst_in_termEnv Ts 1 R) 0 Rm) in
(Term_WT nil nil ts (TermListI Ss)) ->
(Pointwise_subtyping nil Ss Ts1) ->
let Ts' := (typeSubst_in_termEnv (typeSubst_in_termEnv
Ts 1 (typeSubst_in_type Ri0 0 Ri)) 0 Rm) in
let Us := ((Class_type Ci Ri) :: Ts') in
let tb1 := (typeSubst_in_term (typeSubst_in_term tb 1 Ri) 0 Rm) in
let S1 := (typeSubst_in_type (typeSubst_in_type S 1 Ri) 0 Rm) in
(Term_WT nil Us tb1 (TermI S1)) ->
exists S': TypeS,
(Term_WT nil nil (polyTermSubst_in_term tb1 0 (t :: (list_of_termList ts)))
(TermI S')) /\
(Subtyping nil S' S1)
.
Proof
.
intros until PiCi.
intros Hy1 Hy2 Hy3 Hy4 Hy5 Hy6 Hy7 Hy8 Hy9.
intros Ts1 Hy10 Hy11 Ts' Us tb1 S1 Hy12.
set (us:= t :: (list_of_termList ts)).
set (Us':= N :: Ss).
assert (Hy13: Forall2 (fun (u : Term) (U' : TypeS) => Term_WT nil nil u (TermI U'))
us Us').
unfold us. unfold Us'. constructor. trivial. apply Typed_term_list. trivial.
assert (Hy14: Pointwise_subtyping nil Us' Us).
unfold Us. unfold Us'. constructor. trivial.
unfold Ts'.
unfold typeSubst_in_termEnv.
rewrite <- (map_id_is_id _ Ss). apply Forall2_of_map_cor1.
rewrite <- (map_id_is_id _ Ss). apply Forall2_of_map_cor1.
unfold Pointwise_subtyping in Hy11. unfold Ts1 in Hy11.
move Hy11 after Hy13.
apply Forall2_implies with (P1:= (fun T U : TypeS =>
(Subtyping nil T (typeSubst_in_type (typeSubst_in_type U 1 R) 0 Rm)))).
intros. rename xA into T. rename xB into U.
apply S_beta_right_star with
(U':= (typeSubst_in_type (typeSubst_in_type U 1 R') 0 Rm)).
eapply Reduction_preserves_subtyping. apply H. apply RedTypeEnv_refl.
apply RedType_refl.
apply Substitution_in_reduction_star. apply Reduction_in_substitution_star. trivial.
apply Substitution_in_reduction_star. apply Reduction_in_substitution_star. trivial.
apply Forall2_of_map_cor2 with
(P:= fun T U => (Subtyping nil T (typeSubst_in_type U 0 Rm)))
(f:= fun T: TypeS => T)
(g:= fun U => (typeSubst_in_type U 1 R))
.
apply Forall2_of_map_cor2 with
(P:= fun T U => (Subtyping nil T U))
(f:= fun T: TypeS => T)
(g:= fun U => (typeSubst_in_type U 0 Rm))
.
rewrite map_id_is_id. rewrite map_id_is_id.
trivial.
assert (Hy15: TermEnv_WK nil Us).
unfold Us. constructor. inversion Hy4; trivial.
unfold Ts'.
unfold typeSubst_in_termEnv. apply List_forall_of_map.
apply List_forall_implies_forall with
(P1:= fun T => (Type_WK ((typeSubst_in_type Pim 0 R') :: nil) T Type_)).
intros. rename x into T.
inversion Hy5.
eapply Substitution_preserves_well_kindedness_snoc. eauto.
eapply Satisfaction_implies_same_kind; eauto. trivial.
apply List_forall_implies_forall with
(P1:= fun T =>
(Type_WK ((typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri)) :: nil) T Type_)).
intros. rename x into T.
eapply Well_kindedness_is_preserved_by_type_env_reduction_star. eauto.
apply RTE_head_star.
apply Reduction_in_substitution_star. trivial.
apply List_forall_of_map.
change (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri) :: nil) with
(nil ++ (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri) :: nil)).
unfold Pim. unfold Ts. apply Substitution_in_method_param_types_2. trivial.
eapply Poly_term_substitution_preserves_typing_cor2; eauto.
Qed
.
Lemma
Subject_reduction_case_T_call_redex_aux4:
forall (m: MethodSym) (Ci: ClassSym) (R Ri Ri0 R' Rm S N: TypeS) (tb t ts: Term)
(Ss: (list TypeS)),
let C := (methodOwner m) in
let Pim := (methodTypeParams m) in
let Ts := (methodParams m) in
let T := (methodType m) in
let PiCi := (classTypeParams Ci) in
let D := (PiCi :: (typeSubst_in_type Pim 0 Ri0) :: nil) in
let G := ((Class_type Ci (Type_var 1)) :: (typeSubst_in_termEnv Ts 1 Ri0)) in
(Type_WK nil (Class_type C R) Type_) ->
(Type_WK (PiCi :: nil) (Class_type C Ri0) Type_) ->
(Type_WK nil (Class_type C (typeSubst_in_type Ri0 0 Ri)) Type_) ->
(Conform nil Ci Ri) ->
(SectionSat nil Rm (typeSubst_in_type Pim 0 R)) ->
(SectionSat nil Rm (typeSubst_in_type Pim 0 R')) ->
(RedTypeStar R R') ->
(RedTypeStar (typeSubst_in_type Ri0 0 Ri) R') ->
(Type_WK D S Type_) ->
(Subtyping D S (typeSubst_in_type T 1 Ri0)) ->
(Term_WT nil nil t (TermI N)) ->
(Subtyping nil N (Class_type Ci Ri)) ->
let Ts1 := (typeSubst_in_termEnv (typeSubst_in_termEnv Ts 1 R) 0 Rm) in
(Term_WT nil nil ts (TermListI Ss)) ->
(Pointwise_subtyping nil Ss Ts1) ->
(Term_WT D G tb (TermI S)) ->
(TermEnv_WK D G) ->
let tb1 := (typeSubst_in_term (typeSubst_in_term tb 1 Ri) 0 Rm) in
exists S': TypeS,
(Term_WT nil nil (polyTermSubst_in_term tb1 0 (t :: (list_of_termList ts)))
(TermI S')) /\
(Subtyping nil S' (typeSubst_in_type (typeSubst_in_type T 1 R) 0 Rm))
.
Proof
.
intros until G.
intros Hy1 Hy2 Hy3 Hy4 Hy5 Hy6 Hy7 Hy8 Hy9 Hy10 Hy11 Hy12.
intros Ts1 Hy13 Hy14 Hy15 Hy42 tb1.
assert (Hy16: (Type_WK nil (typeSubst_in_type (typeSubst_in_type S 1 Ri) 0 Rm) Type_)).
inversion Hy6.
eapply Type_substitution_preserves_well_kindedness_cor2 with
(Pi:= (typeSubst_in_type Pim 0 R')). eauto.
rewrite <- H. constructor; trivial.
eapply Well_kindedness_is_preserved_by_type_env_reduction_star with
(D:= (typeSubst_in_type Pim 0 (typeSubst_in_type Ri0 0 Ri)) :: nil).
Focus 2. apply RTE_head_star.
apply Reduction_in_substitution_star. trivial.
replace Pim with (typeSubst_in_type Pim 1 Ri).
Focus 2. unfold Pim. apply methodTypeParams_closed.
rewrite <- Permuting_type_substitutions_cor1.
inversion Hy4. inversion H3.
eapply Type_substitution_preserves_well_kindedness_cor1. eauto.
eapply Satisfaction_implies_same_kind. apply H9. apply H4.
trivial.
cut (
exists S': TypeS,
(Term_WT nil nil (polyTermSubst_in_term tb1 0 (t :: (list_of_termList ts)))
(TermI S')) /\
(Subtyping nil S' (typeSubst_in_type (typeSubst_in_type S 1 Ri) 0 Rm))
).
intros. destruct H. destruct H. rename x into S'. rename H into Hy17.
rename H0 into Hy18.
exists S'. split. trivial.
apply S_Trans with (S:= (typeSubst_in_type (typeSubst_in_type S 1 Ri) 0 Rm)).
trivial. trivial.
unfold T.
eapply Subject_reduction_case_T_call_redex_aux1; eauto. apply Hy2.
unfold tb1.
eapply Subject_reduction_case_T_call_redex_aux3.
unfold C in Hy1. apply Hy1. unfold PiCi in Hy2. apply Hy2.
trivial. trivial. apply Hy6. trivial. trivial. apply Hy11. trivial.
eauto. trivial.
eapply Subject_reduction_case_T_call_redex_aux2; eauto.
Qed
.
Theorem
Subject_reduction:
forall (t u: Term) (info: TermInfo),
(Term_WT nil nil t info) ->
(RedTerm t u) ->
exists info': TermInfo,
(Term_WT nil nil u info') /\
(SubInfo nil info' info)
.
End
SetProgram.
Index
This page has been generated by coqdoc