remove module wraps in each file
This commit is contained in:
parent
05c137c489
commit
f53f226f55
9 changed files with 3 additions and 70 deletions
|
@ -1,8 +1,5 @@
|
|||
From Coq Require Import Strings.String.
|
||||
Require Import terms.
|
||||
Include Terms.
|
||||
|
||||
Module Context.
|
||||
|
||||
|
||||
Inductive context : Type :=
|
||||
|
@ -16,6 +13,3 @@ Inductive context_contains : context -> string -> type_term -> Prop :=
|
|||
| C_shuffle : forall x X y Y Γ,
|
||||
(context_contains Γ x X) ->
|
||||
(context_contains (ctx_assign y Y Γ) x X).
|
||||
|
||||
|
||||
End Context.
|
||||
|
|
|
@ -34,12 +34,6 @@ Require Import terms.
|
|||
Require Import subst.
|
||||
From Coq Require Import Strings.String.
|
||||
|
||||
Include Terms.
|
||||
Include Subst.
|
||||
|
||||
Module Equiv.
|
||||
|
||||
|
||||
(** Alpha conversion in types *)
|
||||
|
||||
Reserved Notation "S '--->α' T" (at level 40).
|
||||
|
@ -407,6 +401,3 @@ Proof.
|
|||
apply TEq_Distribute.
|
||||
apply L_DistributeOverSpec2.
|
||||
Qed.
|
||||
|
||||
|
||||
End Equiv.
|
||||
|
|
11
coq/morph.v
11
coq/morph.v
|
@ -5,14 +5,6 @@ Require Import equiv.
|
|||
Require Import subtype.
|
||||
Require Import context.
|
||||
|
||||
Include Terms.
|
||||
Include Subst.
|
||||
Include Equiv.
|
||||
Include Subtype.
|
||||
Include Context.
|
||||
|
||||
Module Morph.
|
||||
|
||||
(* Given a context, there is a morphism path from τ to τ' *)
|
||||
Reserved Notation "Γ '|-' σ '~>' τ" (at level 101, σ at next level, τ at next level).
|
||||
|
||||
|
@ -97,6 +89,3 @@ Proof.
|
|||
apply C_shuffle.
|
||||
apply C_take.
|
||||
Qed.
|
||||
|
||||
|
||||
End Morph.
|
||||
|
|
|
@ -1,14 +1,9 @@
|
|||
From Coq Require Import Strings.String.
|
||||
Require Import terms.
|
||||
Require Import subst.
|
||||
Require Import subtype.
|
||||
Require Import typing.
|
||||
|
||||
Include Terms.
|
||||
Include Subst.
|
||||
Include Typing.
|
||||
|
||||
Module Smallstep.
|
||||
|
||||
Reserved Notation " s '-->α' t " (at level 40).
|
||||
Reserved Notation " s '-->β' t " (at level 40).
|
||||
|
||||
|
@ -140,5 +135,3 @@ Proof.
|
|||
|
||||
apply Multi_Refl.
|
||||
Qed.
|
||||
|
||||
End Smallstep.
|
||||
|
|
|
@ -3,19 +3,11 @@ Require Import terms.
|
|||
Require Import subst.
|
||||
Require Import equiv.
|
||||
Require Import subtype.
|
||||
Require Import context.
|
||||
Require Import morph.
|
||||
Require Import smallstep.
|
||||
Require Import typing.
|
||||
|
||||
Include Terms.
|
||||
Include Subst.
|
||||
Include Equiv.
|
||||
Include Subtype.
|
||||
Include Smallstep.
|
||||
Include Typing.
|
||||
|
||||
|
||||
Module Soundness.
|
||||
|
||||
|
||||
Lemma typing_weakening : forall Γ e τ x σ,
|
||||
(Γ |- e \is τ) ->
|
||||
|
@ -414,5 +406,3 @@ Proof.
|
|||
admit.
|
||||
Admitted.
|
||||
|
||||
End Soundness.
|
||||
|
||||
|
|
|
@ -1,10 +1,6 @@
|
|||
From Coq Require Import Strings.String.
|
||||
Require Import terms.
|
||||
|
||||
Include Terms.
|
||||
|
||||
Module Subst.
|
||||
|
||||
(* Type Variable "x" is a free variable in type *)
|
||||
Inductive type_var_free (x:string) : type_term -> Prop :=
|
||||
| TFree_Var :
|
||||
|
@ -128,5 +124,3 @@ Fixpoint expr_specialize (v:string) (n:type_term) (e0:expr_term) :=
|
|||
| expr_descend t e => expr_descend (type_subst v n t) (expr_specialize v n e)
|
||||
| e => e
|
||||
end.
|
||||
|
||||
End Subst.
|
||||
|
|
|
@ -11,10 +11,6 @@
|
|||
From Coq Require Import Strings.String.
|
||||
Require Import terms.
|
||||
Require Import equiv.
|
||||
Include Terms.
|
||||
Include Equiv.
|
||||
|
||||
Module Subtype.
|
||||
|
||||
(** Subtyping *)
|
||||
|
||||
|
@ -96,5 +92,3 @@ Proof.
|
|||
apply TSubRepr_Refl.
|
||||
apply TEq_Refl.
|
||||
Qed.
|
||||
|
||||
End Subtype.
|
||||
|
|
|
@ -4,7 +4,6 @@
|
|||
*)
|
||||
|
||||
From Coq Require Import Strings.String.
|
||||
Module Terms.
|
||||
|
||||
(* types *)
|
||||
Inductive type_term : Type :=
|
||||
|
@ -118,5 +117,3 @@ Compute polymorphic_identity1.
|
|||
|
||||
Close Scope ladder_type_scope.
|
||||
Close Scope ladder_expr_scope.
|
||||
|
||||
End Terms.
|
||||
|
|
|
@ -9,15 +9,6 @@ Require Import subtype.
|
|||
Require Import context.
|
||||
Require Import morph.
|
||||
|
||||
Include Terms.
|
||||
Include Subst.
|
||||
Include Equiv.
|
||||
Include Subtype.
|
||||
Include Context.
|
||||
Include Morph.
|
||||
|
||||
Module Typing.
|
||||
|
||||
(** Typing Derivation *)
|
||||
|
||||
Reserved Notation "Gamma '|-' x '\is' X" (at level 101, x at next level, X at level 0).
|
||||
|
|
Loading…
Reference in a new issue