-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathlambda2.v
130 lines (103 loc) · 3.99 KB
/
lambda2.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
Import Nat.
Require Import List.
Import List.ListNotations.
Definition atomic_type := nat.
Inductive type : Type :=
| Var : atomic_type -> type
| Arr : type -> type -> type
| Pi : atomic_type -> type -> type.
Notation "& x" := (Var x) (at level 8).
Infix ">>" := Arr (right associativity, at level 9).
Fixpoint FVb (n : atomic_type) (t : type) : bool :=
match t with
| &m => eqb n m
| a >> b => (FVb n a) || (FVb n b)
| Pi m t => (negb (eqb m n)) && (FVb n t)
end.
Fixpoint FVl (t : type) : list atomic_type :=
match t with
| &m => [m]
| a >> b => (FVl a) ++ (FVl b)
| Pi m t => remove PeanoNat.Nat.eq_dec m (FVl t)
end.
Definition atomic_term := nat.
Inductive term : Type :=
| TVar : atomic_term -> term
| App : term -> term -> term
| Tapp : term -> type -> term
| Abs : atomic_term -> type -> term -> term
| Tabs : atomic_type -> term -> term.
Notation "# x" := (TVar x) (at level 10).
Infix "!" := (App) (left associativity, at level 11).
Infix "!!" := (Tapp) (left associativity, at level 12).
Notation "\ x t m" := (Abs x t m) (right associativity, at level 13).
Notation "/ t m" := (Tabs t m) (right associativity, at level 14).
Inductive statement : Type :=
| St : term -> type -> statement
| TSt : type -> statement.
Infix ":#" := St (at level 10).
Inductive declaration : Type :=
| Std : atomic_term -> type -> declaration
| TStd : atomic_type -> declaration.
Definition dec2st (d : declaration) : statement :=
match d with
| Std n a => St (TVar n) a
| TStd n => TSt (Var n)
end.
Infix ":*" := Std (at level 10).
Definition context := list declaration.
Fixpoint check_type (G : context) (t : atomic_type) : bool :=
match G with
| nil => false
| cons d G' =>
match d with
| _ :* _ => check_type G' t
| TStd n => eqb n t || check_type G' t
end
end.
Fixpoint check_term (G : context) (x : atomic_term) : bool :=
match G with
| nil => false
| cons d G' =>
match d with
| y :* _ => eqb x y || check_term G' x
| TStd _ => check_term G' x
end
end.
Definition foreach {T : Type} (l : list T) (P : T -> bool) := fold_left andb (map P l) true = true.
Coercion atomic_type_as_declaration (a : atomic_type) : declaration := TStd a.
Inductive l2_context : context -> Prop :=
| Emp : l2_context []
| CType : forall (G : context) (a : atomic_type), l2_context G -> (check_type G a = false) -> l2_context ((TStd a) :: G)
| CTerm : forall (G : context) (x : atomic_term) (r : type), l2_context G -> (check_term G x = false) -> (foreach (FVl r) (check_type G)) -> l2_context ((x :* r) :: G).
Fixpoint subs (t : type) (a : atomic_type) (b : type) : type :=
match t with
| & x => if eqb x a then b else t
| r >> s => (subs r a b) >> (subs r a b)
| Pi x s => if eqb x a then t else Pi x (subs s a b)
end.
Coercion atomic_term_as_term (a : atomic_term) : term := # a.
Reserved Notation "G ⊢ st" (no associativity, at level 61). (* 22a2 *)
Inductive l2 : context -> statement -> Prop :=
| Var2 : forall (G : context) (x : atomic_term) (s : type), (l2_context G) -> (In (x :* s) G) -> G ⊢ x :# s
| App2 : forall G M N s t, (G ⊢ (M :# (s >> t))) -> (G ⊢ (N :# s)) -> G ⊢ ((M ! N ) :# t)
| Abs2 : forall G M x s t, ((x :* s) :: G ⊢ M :# t) -> G ⊢ (Abs x s M) :# (s >> t)
| Form2 : forall G B, (l2_context G) -> (foreach (FVl B) (check_type G)) -> G ⊢ TSt B
| Tapp2 : forall G M a A B, (G ⊢ (M :# (Pi a A))) -> (G ⊢ (TSt B)) -> G ⊢ ((M !! B) :# (subs A a B))
| Tabs2 : forall G a M A, ((TStd a)::G ⊢ M :# A) -> G ⊢ (Tabs a M) :# (Pi a A)
where "G ⊢ st" := (l2 G st).
Lemma context_term : forall G x s, l2_context ((x :* s) :: G) -> l2_context G.
Proof. intros G x s H. inversion H. auto. Qed.
Lemma context_type : forall G a, l2_context ((TStd a) :: G) -> l2_context G.
Proof. intros G a H. inversion H. auto. Qed.
Theorem exercise19: forall G l s, G ⊢ l :# s -> l2_context G.
Proof.
intros.
induction H.
- apply H.
- apply IHl2_1.
- apply (context_term _ _ _ IHl2).
- apply H.
- apply IHl2_1.
- apply (context_type _ _ IHl2).
Qed.