├── .gitignore ├── TyDE-2019 └── Names.v ├── examples ├── Lambda.v ├── Makefile └── _CoqProject └── src ├── Context.v ├── Makefile ├── Morph.v ├── Relative.v ├── Var.v └── _CoqProject /.gitignore: -------------------------------------------------------------------------------- 1 | ### Coq ### 2 | .*.aux 3 | *.a 4 | *.cma 5 | *.cmi 6 | *.cmo 7 | *.cmx 8 | *.cmxa 9 | *.cmxs 10 | *.glob 11 | *.ml.d 12 | *.ml4.d 13 | *.mli.d 14 | *.mllib.d 15 | *.mlpack.d 16 | *.native 17 | *.o 18 | *.v.d 19 | *.vio 20 | *.vo 21 | .coq-native/ 22 | .csdp.cache 23 | .lia.cache 24 | .nia.cache 25 | .nlia.cache 26 | .nra.cache 27 | csdp.cache 28 | lia.cache 29 | nia.cache 30 | nlia.cache 31 | nra.cache -------------------------------------------------------------------------------- /TyDE-2019/Names.v: -------------------------------------------------------------------------------- 1 | Require Import String. 2 | 3 | Definition Zero := Empty_set. 4 | Inductive Succ {S : Set} : Set := v0 | vS (s : S). 5 | 6 | Fixpoint level (V : nat) : Set := 7 | match V with 8 | | 0 => Zero 9 | | S V => @Succ (level V) 10 | end. 11 | 12 | Inductive var {V : nat} := 13 | | free (name : string) (index : nat) 14 | | bound (v : level V). 15 | 16 | Definition wkv {V} (v : @var V) : @var (S V) := 17 | match v with 18 | | free n i => free n i 19 | | bound v => @bound (S V) (vS v) 20 | end. 21 | 22 | Definition openv a {V} (v : @var (S V)) : @var V := 23 | match v with 24 | | free xn xi => 25 | if string_dec a xn then 26 | free xn (S xi) 27 | else 28 | free xn xi 29 | | bound (vS v) => bound v 30 | | bound v0 => free a 0 31 | end. 32 | 33 | Definition closev a {V} (v : @var V) : @var (S V) := 34 | match v with 35 | | free xn xi => 36 | if string_dec a xn then 37 | match xi with 38 | | 0 => @bound (S V) v0 39 | | S xi' => free xn xi' 40 | end 41 | else 42 | free xn xi 43 | | bound v => @bound (S V) (vS v) 44 | end. 45 | 46 | Definition bindv {V} (v : @var (S V)) : option (@var V) := 47 | match v with 48 | | free xn xi => Some (free xn xi) 49 | | bound v0 => None 50 | | bound (vS v) => Some (@bound V v) 51 | end. 52 | 53 | Notation shiftv a v := (openv a (wkv v)). 54 | 55 | Arguments openv : simpl nomatch. 56 | Arguments closev : simpl nomatch. 57 | Arguments bindv : simpl nomatch. 58 | 59 | Ltac string_eqs := 60 | match goal with 61 | | |- context C [string_dec ?x ?x] => 62 | destruct (string_dec x x) as [_ | Neq]; cbn; try easy 63 | | |- context C [string_dec ?x ?y] => 64 | let Neq := fresh "Hneq_" x y in 65 | destruct (string_dec x y) as [-> | Neq]; cbn; try easy 66 | end. 67 | 68 | Lemma rw_openv_closev a {V} (v : @var V) : openv a (closev a v) = v. 69 | Proof. 70 | unfold openv, closev. 71 | destruct v as [vn [|vi]|v]; cbn; try easy; repeat string_eqs. 72 | Qed. 73 | Lemma rw_closev_same a {V} : 74 | @closev a V (free a 0) = @bound (S V) v0. 75 | Proof. unfold closev. string_eqs. Qed. 76 | Lemma rw_closev_openv a {V} (v : @var (S V)) : closev a (openv a v) = v. 77 | Proof. 78 | unfold openv, closev. 79 | destruct v as [vn vi|[|v]]; cbn; try easy; repeat string_eqs. 80 | Qed. 81 | Lemma rw_bindv_wkv {V} (v : @var V) : bindv (wkv v) = Some v. 82 | Proof. destruct v; easy. Qed. 83 | Lemma rw_shiftv_wkv b {V} (v : @var V) : 84 | shiftv b (wkv v) = wkv (shiftv b v). 85 | Proof. unfold openv; destruct v; cbn; try easy. string_eqs. Qed. 86 | Hint Rewrite @rw_closev_openv @rw_closev_same @rw_openv_closev @rw_bindv_wkv 87 | @rw_shiftv_wkv : rw_names. 88 | (* Slightly weird, but useful *) 89 | Lemma rw_closev_wkv_shiftv a {V} (v : @var V) : 90 | closev a (wkv (shiftv a v)) = wkv (wkv v). 91 | Proof. rewrite <- rw_shiftv_wkv. rewrite rw_closev_openv. easy. Qed. 92 | Hint Rewrite @rw_closev_wkv_shiftv : rw_names. 93 | 94 | Lemma compare_aux a {V} (x : @var V) : 95 | (x = free a 0) + { xi' | x = shiftv a xi' }. 96 | Proof. 97 | unfold openv. 98 | destruct x as [xn xi | x]; cbn. 99 | destruct (string_dec a xn) as [->|Neq]. 100 | destruct xi. 101 | - left; easy. 102 | - right; exists (free xn xi); cbn. string_eqs. 103 | - right; exists (free xn xi); cbn. string_eqs. 104 | - right; exists (bound x); easy. 105 | Qed. 106 | 107 | Inductive var_comparison (a : string) {V} : @var V -> Set := 108 | | samev : var_comparison a (@free V a 0) 109 | | diffv x : var_comparison a (shiftv a x). 110 | Definition compare a {V} (x : @var V) : var_comparison a x. 111 | destruct (compare_aux a x) as [-> | [xi ->]]; constructor. Defined. 112 | 113 | 114 | Inductive v0_comparison {V} : @var (S V) -> Set := 115 | | samev0 : v0_comparison (@bound (S V) v0) 116 | | diffv0 v : v0_comparison (wkv v). 117 | Definition compare0 {V} (v : @var (S V)) : v0_comparison v. 118 | Proof. 119 | destruct v as [v i|[|v]]. 120 | - change (@free (S V) v i) with (@wkv V (free v i)); constructor. 121 | - constructor. 122 | - change (@bound (S V) (vS v)) with (@wkv V (bound v)); constructor. 123 | Qed. 124 | 125 | 126 | Inductive term {V : nat} := 127 | | vv : @var V -> @term V 128 | | const : nat -> @term V 129 | | lam : @term (S V) -> @term V 130 | | app : @term V -> @term V -> @term V. 131 | 132 | Fixpoint open a {V} (t : @term (S V)) := 133 | match t with 134 | | vv v => vv (openv a v) 135 | | const k => const k 136 | | lam t => lam (open a t) 137 | | app f e => app (open a f) (open a e) 138 | end. 139 | 140 | Fixpoint close a {V} (t : @term V) := 141 | match t with 142 | | vv v => vv (closev a v) 143 | | const k => const k 144 | | lam t => lam (close a t) 145 | | app f e => app (close a f) (close a e) 146 | end. 147 | 148 | Fixpoint wk {V} (t : @term V) := 149 | match t with 150 | | vv v => vv (wkv v) 151 | | const k => const k 152 | | lam t => lam (wk t) 153 | | app f e => app (wk f) (wk e) 154 | end. 155 | 156 | Fixpoint wkn {V} (t : @term 0) := 157 | match V with 158 | | 0 => t 159 | | S V => wk (wkn t) 160 | end. 161 | 162 | Fixpoint bind {V} (u : @term 0) (t : @term (S V)) := 163 | match t with 164 | | vv v => 165 | match bindv v with 166 | | Some v => vv v 167 | | None => wkn u 168 | end 169 | | const k => const k 170 | | lam t => lam (bind u t) 171 | | app f e => app (bind u f) (bind u e) 172 | end. 173 | 174 | 175 | Inductive ren := 176 | | r_id 177 | | r_comp (r : ren) (s : ren) 178 | | r_shift (b : string) (r : ren) 179 | | r_rename (b : string) (r : ren) (a : string) 180 | | r_subst (t : @term 0) (r : ren) (a : string). 181 | Notation r_under a r := (r_rename a r a). 182 | 183 | Fixpoint map_term (f : forall V, @var V -> @term V) {V} (t : @term V) : @term V := 184 | match t with 185 | | vv v => f _ v 186 | | const k => const k 187 | | lam t => lam (map_term f t) 188 | | app a b => app (map_term f a) (map_term f b) 189 | end. 190 | 191 | 192 | Fixpoint applyv (r : ren) {V} (v : @var V) {struct r} : @term V := 193 | match r with 194 | | r_id => vv v 195 | | r_comp r s => map_term (@applyv r) (applyv s v) 196 | | r_shift b r => open b (applyv r (wkv v)) 197 | | r_rename b r a => open b (applyv r (closev a v)) 198 | | r_subst u r a => bind u (applyv r (closev a v)) 199 | end. 200 | Arguments applyv !r V v /. 201 | Notation apply r := (map_term (@applyv r)). 202 | 203 | Fixpoint is_renaming (r : ren) : Prop := 204 | match r with 205 | | r_id => True 206 | | r_comp r s => is_renaming r /\ is_renaming s 207 | | r_shift b r => is_renaming r 208 | | r_rename b r a => is_renaming r 209 | | r_subst u r a => False 210 | end. 211 | 212 | Definition proj1 {A B : Prop} (H : A /\ B) := let (a, _) := H in a. 213 | Definition proj2 {A B : Prop} (H : A /\ B) := let (_, b) := H in b. 214 | Fixpoint applyrv r : forall (rn : is_renaming r) {V} (v : @var V), @var V := 215 | match r with 216 | | r_id => fun _ _ v => v 217 | | r_comp r s => fun rn _ v => applyrv r (proj1 rn) (applyrv s (proj2 rn) v) 218 | | r_shift b r => fun rn _ v => openv b (applyrv r rn (wkv v)) 219 | | r_rename b r a => fun rn _ v => openv b (applyrv r rn (closev a v)) 220 | | r_subst _ _ _ => False_rec _ 221 | end. 222 | 223 | Definition apply' r {V} (t : @term V) := 224 | match r with 225 | | r_id => t 226 | | r_comp r s => apply r (apply s t) 227 | | r_shift b r => open b (apply r (wk t)) 228 | | r_rename b r a => open b (apply r (close a t)) 229 | | r_subst u r a => bind u (apply r (close a t)) 230 | end. 231 | Arguments apply' !r V t /. 232 | Lemma expand_apply r {V} {t : @term V} : 233 | apply r t = apply' r t. 234 | Proof. induction t; destruct r; cbn; f_equal; auto. Qed. 235 | Ltac expand_apply := 236 | repeat (rewrite expand_apply; cbn); repeat rewrite <- expand_apply. 237 | 238 | (* Better dependent induction *) 239 | Definition induction_S 240 | (ty : nat -> Set) 241 | (P : forall V, ty (S V) -> Prop) 242 | (Thm : forall V (t : ty V), 243 | match V return ty V -> Prop with 244 | | 0 => fun _ => True 245 | | S V => P V 246 | end t) : forall V t, P V t := 247 | fun V t => Thm (S V) t. 248 | 249 | Ltac induction_S := 250 | match goal with 251 | | |- forall (V : nat) (tm : ?ty (S V)), @?P V tm => 252 | apply (@induction_S ty P); 253 | induction tm; 254 | match goal with 255 | | |- match ?V with 0 => fun _ => True | S _ => _ end _ => 256 | destruct V; [exact I | idtac] 257 | | _ => idtac 258 | end 259 | end. 260 | 261 | Lemma rw_close_open a : 262 | forall {V} (t : @term (S V)), close a (open a t) = t. 263 | Proof. 264 | induction_S; cbn; try (f_equal; easy). 265 | autorewrite with rw_names; easy. 266 | Qed. 267 | 268 | Lemma rw_open_close a : 269 | forall {V} (t : @term V), open a (close a t) = t. 270 | Proof. 271 | induction t; cbn; try (f_equal; easy). 272 | autorewrite with rw_names; easy. 273 | Qed. 274 | 275 | Lemma rw_bind_wk u : 276 | forall {V} (t : @term V), bind u (wk t) = t. 277 | Proof. 278 | induction t; cbn; try (f_equal; easy). 279 | autorewrite with rw_names; easy. 280 | Qed. 281 | 282 | Hint Rewrite @rw_close_open @rw_open_close @rw_bind_wk : rw_names. 283 | 284 | 285 | Notation shift a t := (apply (r_shift a r_id) t). 286 | 287 | (* identity and composition for apply *) 288 | Lemma rw_apply_id {V} (t : @term V) : apply r_id t = t. 289 | Proof. rewrite expand_apply; easy. Qed. 290 | Lemma rw_apply_comp r s {V} (t : @term V) : 291 | apply r (apply s t) = apply (r_comp r s) t. 292 | Proof. rewrite (expand_apply (r_comp r s)); easy. Qed. 293 | Hint Rewrite @rw_apply_id @rw_apply_comp : rw_names. 294 | 295 | (* Simple rewritings on applyv and shiftv (more later) *) 296 | Lemma rw_applyv_bound r : forall {V} (v : level V), 297 | applyv r (bound v) = vv (bound v). 298 | Proof. 299 | induction r; cbn; intros; autorewrite with rw_names; cbn; 300 | repeat match goal with [ H : _ |- _ ] => rewrite H; clear H end; 301 | cbn; easy. 302 | Qed. 303 | Lemma rw_applyrv_bound r (rn : is_renaming r) : forall {V} (v : level V), 304 | applyrv r rn (bound v) = bound v. 305 | Proof. 306 | induction r; try destruct rn; cbn; intros; auto; 307 | repeat match goal with [ H : _ |- _ ] => rewrite H; clear H end; 308 | cbn; easy. 309 | Qed. 310 | Hint Rewrite @rw_applyv_bound (*@rw_shift_shiftv*) @rw_applyrv_bound: rw_names. 311 | 312 | (* Rewrite to do the following: 313 | - group balanced pairs of operations into 'ren's 314 | - push opens/bind right and wk/close left 315 | - simplify rens *) 316 | 317 | Lemma rw_group_rename a b {V} (t : @term V) : 318 | open a (close b t) = apply (r_rename a r_id b) t. 319 | Proof. expand_apply; easy. Qed. 320 | Lemma rw_group_shift a {V} (t : @term V) : 321 | open a (wk t) = apply (r_shift a r_id) t. 322 | Proof. expand_apply; easy. Qed. 323 | Lemma rw_group_subst u a {V} (t : @term V) : 324 | bind u (close a t) = apply (r_subst u r_id a) t. 325 | Proof. expand_apply; easy. Qed. 326 | Hint Rewrite @rw_group_rename @rw_group_shift @rw_group_subst : rw_names. 327 | 328 | 329 | (* wk commutes with shift *) 330 | Lemma rw_shift_wk b {V} (t : @term V) : 331 | shift b (wk t) = wk (shift b t). 332 | Proof. 333 | induction t; cbn; try (f_equal; auto). 334 | autorewrite with rw_names; easy. 335 | Qed. 336 | Hint Rewrite @rw_shift_wk : rw_names. 337 | 338 | (* wk commutes with apply. 339 | Somewhat harder to prove than I'd expect *) 340 | 341 | Lemma apply_if_applyv_wk r : 342 | (forall {V} (v : @var V), applyv r (wkv v) = wk (applyv r v)) -> 343 | (forall {V} (t : @term V), apply r (wk t) = wk (apply r t)). 344 | Proof. intro H. induction t; cbn; auto; f_equal; auto. Qed. 345 | 346 | Lemma rw_applyv_wkv r : forall {V} (v : @var V), applyv r (wkv v) = wk (applyv r v). 347 | Proof. 348 | induction r; cbn; intros. 349 | - easy. 350 | - rewrite IHr2. apply apply_if_applyv_wk. auto. 351 | 352 | - repeat rewrite IHr. 353 | autorewrite with rw_names. easy. 354 | 355 | - destruct (compare a v); cbn; 356 | autorewrite with rw_names; cbn; autorewrite with rw_names; cbn; try easy. 357 | repeat rewrite IHr. 358 | autorewrite with rw_names. easy. 359 | 360 | - destruct (compare a v); cbn; 361 | autorewrite with rw_names; cbn; autorewrite with rw_names; cbn; try easy. 362 | repeat rewrite IHr. 363 | autorewrite with rw_names; easy. 364 | Qed. 365 | 366 | 367 | (* FIXME where? *) 368 | Lemma applyrv_is_applyv : 369 | forall r (rn : is_renaming r) {V} (v : @var V), 370 | applyv r v = vv (applyrv r rn v). 371 | Proof. 372 | induction r; cbn; intuition. 373 | - erewrite IHr2; cbn. erewrite IHr1. reflexivity. 374 | - erewrite IHr. reflexivity. 375 | - erewrite IHr. cbn. reflexivity. 376 | Qed. 377 | 378 | Lemma rw_applyrv_wkv : 379 | forall r (rn : is_renaming r) {V} (v : @var V), 380 | applyrv r rn (wkv v) = wkv (applyrv r rn v). 381 | Proof. 382 | intros. 383 | assert (Inj : forall {V} (a b : @var V), vv a = vv b -> a = b) by congruence. 384 | apply Inj. rewrite <- applyrv_is_applyv. 385 | rewrite rw_applyv_wkv. 386 | rewrite (applyrv_is_applyv r rn). 387 | easy. 388 | Qed. 389 | 390 | Lemma rw_apply_wk r {V} (t : @term V) : apply r (wk t) = wk (apply r t). 391 | Proof. apply apply_if_applyv_wk. apply rw_applyv_wkv. Qed. 392 | 393 | Lemma rw_apply_wkn r {V} (t : @term 0) : 394 | apply r (@wkn V t) = @wkn V (apply r t). 395 | Proof. 396 | induction V; cbn; try easy. 397 | rewrite rw_apply_wk; rewrite IHV; easy. 398 | Qed. 399 | 400 | Hint Rewrite @rw_applyv_wkv @rw_applyrv_wkv @rw_apply_wk @rw_apply_wkn : rw_names. 401 | 402 | Lemma rw_applyv_wkv_free r {an ai V} : 403 | applyv r (@free (S V) an ai) = wk (applyv r (@free V an ai)). 404 | Proof. 405 | change (@free (S V) an ai) with (wkv (@free V an ai)). 406 | apply rw_applyv_wkv. 407 | Qed. 408 | Lemma rw_applyrv_wkv_free r rn {an ai V} : 409 | applyrv r rn (@free (S V) an ai) = wkv (applyrv r rn (@free V an ai)). 410 | Proof. 411 | change (@free (S V) an ai) with (wkv (@free V an ai)). 412 | apply rw_applyrv_wkv. 413 | Qed. 414 | Hint Rewrite @rw_applyv_wkv_free @rw_applyrv_wkv_free : rw_names. 415 | 416 | Ltac names := 417 | cbn; 418 | repeat progress (autorewrite with rw_names; cbn). 419 | 420 | (* Reductions for applyv to particular variables *) 421 | Lemma rw_applyv_shift_comp r b {V} (x : @var V) : 422 | applyv (r_comp r (r_shift b r_id)) x = applyv r (shiftv b x). 423 | Proof. easy. Qed. 424 | Lemma rw_applyv_rename_same b r a {V} : 425 | applyv (r_rename b r a) (@free V a 0) = vv (free b 0). 426 | Proof. names; easy. Qed. 427 | Lemma rw_applyv_rename_same_comp r' b r a {V} : 428 | applyv (r_comp r' (r_rename b r a)) (@free V a 0) = applyv r' (free b 0). 429 | Proof. names; easy. Qed. 430 | Lemma rw_applyv_rename_diffv b r a {V} (x : @var V) : 431 | applyv (r_rename b r a) (shiftv a x) = 432 | applyv (r_shift b r) x. 433 | Proof. names; easy. Qed. 434 | Lemma rw_applyv_rename_diffv_comp r' b r a {V} (x : @var V) : 435 | applyv (r_comp r' (r_rename b r a)) (shiftv a x) = 436 | applyv (r_comp r' (r_shift b r)) x. 437 | Proof. names; easy. Qed. 438 | Lemma rw_applyv_subst_same u r a {V} : 439 | applyv (r_subst u r a) (@free V a 0) = wkn u. 440 | Proof. names; easy. Qed. 441 | Lemma rw_applyv_subst_same_comp r' u r a {V} : 442 | applyv (r_comp r' (r_subst u r a)) (@free V a 0) = 443 | wkn (apply r' u). 444 | Proof. names; easy. Qed. 445 | Lemma rw_applyv_subst_diffv u r a {V} (x : @var V) : 446 | applyv (r_subst u r a) (shiftv a x) = applyv r x. 447 | Proof. names; easy. Qed. 448 | Lemma rw_applyv_subst_diffv_comp r' u r a {V} (x : @var V) : 449 | applyv (r_comp r' (r_subst u r a)) (shiftv a x) = 450 | applyv (r_comp r' r) x. 451 | Proof. names; easy. Qed. 452 | Goal forall a b {V} (v : @var (S V)), 453 | vv (openv a (shiftv b v)) = applyv (r_under a (r_shift b r_id)) (openv a v). 454 | intros. cbn. names. easy. Qed. 455 | Goal forall r a {V} (v : @var (S V)), 456 | applyv r (closev a v) = close a (applyv (r_under a r) v). 457 | Proof. intros; names. easy. Qed. 458 | 459 | Lemma comm_bind_apply u : forall {V} (t : @term (S V)) r, 460 | apply r (bind u t) = bind (apply r u) (apply r t). 461 | Proof. 462 | induction_S; intro r; cbn; try solve [f_equal; auto]. 463 | destruct (compare0 v); cbn; autorewrite with rw_names; easy. 464 | Qed. 465 | 466 | 467 | (* FIXME: maybe this should just be "names"? *) 468 | Lemma rw_close_shift a r {V} (t : @term V) : 469 | close a (apply (r_shift a r) t) = wk (apply r t). 470 | Proof. expand_apply. rewrite rw_apply_wk. rewrite rw_close_open. easy. Qed. 471 | Lemma rw_subst_open u r b {V} (t : @term (S V)) : 472 | apply (r_subst u r b) (open b t) = bind u (apply r t). 473 | Proof. expand_apply. rewrite rw_close_open. easy. Qed. 474 | Lemma rw_rename_open a b {V} (t : @term (S V)) : 475 | apply (r_rename a r_id b) (open b t) = open a t. 476 | Proof. expand_apply. rewrite rw_close_open. easy. Qed. 477 | Hint Rewrite @rw_close_shift @rw_subst_open @rw_rename_open : rw_names. 478 | 479 | 480 | (* FIXME: lots of applyrv rewritings below here *) 481 | 482 | 483 | 484 | (* FIXME: not great, tbh *) 485 | 486 | (* Commuting bind with apply. 487 | It seems more natural to make 488 | bind (r u) . r --> r . bind u 489 | since bind is similar to open, and open moves right. 490 | But that rule breaks confluence, and isn't very useful 491 | as bind doesn't show up in terms. So, binds move left 492 | instead. *) 493 | Lemma rw_push_bind u r {V} (t : @term (S V)) : 494 | apply r (bind u t) = bind (apply r u) (apply r t). 495 | Proof. apply comm_bind_apply. Qed. 496 | (* Shifts go the other way. 497 | This is confluent/terminating because shift doesn't simplify to apply (?) *) 498 | Lemma rw_push_bind_shift u r s a {V} (t : @term (S V)) : 499 | bind (apply (r_shift a r) u) (apply (r_shift a s) t) = 500 | shift a (bind (apply r u) (apply s t)). 501 | Proof. rewrite comm_bind_apply. expand_apply. names. easy. Qed. 502 | Lemma rw_push_bind_shiftvu v a {V} (t : @term (S V)) : 503 | bind (vv (shiftv a v)) (shift a t) = shift a (bind (vv v) t). 504 | Proof. 505 | change (vv (shiftv a v)) with (shift a (vv v)). 506 | rewrite rw_push_bind_shift. autorewrite with rw_names. easy. 507 | Qed. 508 | Hint Rewrite 509 | @rw_push_bind @rw_push_bind_shift @rw_push_bind_shiftvu : rw_names. 510 | 511 | 512 | (* Morphisms. *) 513 | Require Import Setoid. 514 | Require Import Morphisms. 515 | 516 | (* FIXME: is it better to define eqr over all vars or just free ones? *) 517 | Definition eqr (r s : ren) : Prop := 518 | forall {V} (v : @var V), applyv r v = applyv s v. 519 | 520 | Instance eqr_Equivalence : Equivalence eqr. 521 | Proof. 522 | split; red; unfold eqr; auto. 523 | intros. etransitivity; eauto. 524 | Qed. 525 | 526 | Instance mor_applyv : Proper 527 | (eqr ==> forall_relation (fun _ => (eq ==> eq))) applyv. 528 | Proof. intros r s Hrs V v v' <-. auto. Qed. 529 | 530 | Instance mor_map_term : Proper 531 | (forall_relation (fun _ => (eq ==> eq)) ==> forall_relation (fun _ => (eq ==> eq))) map_term. 532 | Proof. 533 | intros f g Hfg V t t' <-. cbv in Hfg. 534 | induction t; cbn; try (f_equal; auto). 535 | Qed. 536 | 537 | (* FIXME applyv, applyrv, shift below *) 538 | 539 | 540 | 541 | 542 | Ltac easy_eqr := 543 | repeat intro; cbn; 544 | repeat (match goal with [ H : _ |- _ ] => rewrite H; clear H end); 545 | expand_apply; names; easy. 546 | 547 | Instance mor_r_comp : Proper (eqr ==> eqr ==> eqr) r_comp. easy_eqr. Qed. 548 | Instance mor_r_shift : Proper (eq ==> eqr ==> eqr) r_shift. easy_eqr. Qed. 549 | Instance mor_r_rename : Proper (eq ==> eqr ==> eq ==> eqr) r_rename. easy_eqr. Qed. 550 | Instance mor_r_subst : Proper (eq ==> eqr ==> eq ==> eqr) r_subst. easy_eqr. Qed. 551 | 552 | Lemma rw_ren_id_comp r : eqr (r_comp r_id r) r. 553 | Proof. easy_eqr. Qed. 554 | Lemma rw_ren_comp_id r : eqr (r_comp r r_id) r. 555 | Proof. easy_eqr. Qed. 556 | Lemma rw_ren_comp_assoc r s t : 557 | eqr (r_comp r (r_comp s t)) (r_comp (r_comp r s) t). 558 | Proof. easy_eqr. Qed. 559 | Lemma rw_ren_under_id x : 560 | eqr (r_under x r_id) r_id. 561 | Proof. easy_eqr. Qed. 562 | Lemma rw_ren_rename_rename c b a r s : 563 | eqr (r_comp (r_rename c r b) (r_rename b s a)) 564 | (r_rename c (r_comp r s) a). 565 | Proof. easy_eqr. Qed. 566 | Lemma rw_ren_rename_rename_comp l c b a r s : 567 | eqr (r_comp (r_comp l (r_rename c r b)) (r_rename b s a)) 568 | (r_comp l (r_rename c (r_comp r s) a)). 569 | Proof. easy_eqr. Qed. 570 | Lemma rw_ren_rename_shift c b r s : 571 | eqr (r_comp (r_rename c r b) (r_shift b s)) 572 | (r_shift c (r_comp r s)). 573 | Proof. 574 | repeat intro; cbn. expand_apply. names. easy. Qed. 575 | Lemma rw_ren_rename_shift_comp l c b r s : 576 | eqr (r_comp (r_comp l (r_rename c r b)) (r_shift b s)) 577 | (r_comp l (r_shift c (r_comp r s))). 578 | Proof. 579 | repeat intro; cbn. expand_apply. rewrite rw_close_open; easy. Qed. 580 | Lemma rw_ren_subst_rename u r a s b : 581 | eqr (r_comp (r_subst u r a) (r_rename a s b)) 582 | (r_subst u (r_comp r s) b). 583 | Proof. repeat intro; cbn; expand_apply; repeat rewrite rw_close_open; easy. Qed. 584 | Lemma rw_ren_subst_rename_comp l u r a s b : 585 | eqr (r_comp (r_comp l (r_subst u r a)) (r_rename a s b)) 586 | (r_comp l (r_subst u (r_comp r s) b)). 587 | Proof. repeat intro; cbn; expand_apply; repeat rewrite rw_close_open; easy. Qed. 588 | Lemma rw_ren_subst_shift u r a s : 589 | eqr (r_comp (r_subst u r a) (r_shift a s)) 590 | (r_comp r s). 591 | Proof. 592 | repeat intro; expand_apply; cbn; expand_apply. 593 | rewrite rw_close_open. rewrite rw_applyv_wkv. rewrite rw_apply_wk. rewrite rw_bind_wk. 594 | easy. 595 | Qed. 596 | Lemma rw_ren_subst_shift_comp l u r a s : 597 | eqr (r_comp (r_comp l (r_subst u r a)) (r_shift a s)) 598 | (r_comp l (r_comp r s)). 599 | Proof. 600 | repeat intro; expand_apply; cbn; expand_apply. 601 | rewrite rw_close_open. autorewrite with rw_names. easy. 602 | Qed. 603 | Lemma rw_ren_shift_any a r s : 604 | eqr (r_comp (r_shift a r) s) (r_shift a (r_comp r s)). 605 | Proof. 606 | repeat intro; expand_apply; cbn; expand_apply. 607 | autorewrite with rw_names. easy. 608 | Qed. 609 | Lemma rw_ren_shift_any_comp l a r s : 610 | eqr (r_comp (r_comp l (r_shift a r)) s) (r_comp l (r_shift a (r_comp r s))). 611 | Proof. 612 | repeat intro; expand_apply; cbn; expand_apply. 613 | autorewrite with rw_names. easy. 614 | Qed. 615 | Lemma rw_ren_any_subst r u s x : 616 | eqr (r_comp r (r_subst u s x)) (r_subst (apply r u) (r_comp r s) x). 617 | Proof. easy_eqr. Qed. 618 | 619 | (* No _comp version required for any_subst *) 620 | (* Ensure crit pair of rw_ren_shift_any and rw_ren_any_subst is confluent *) 621 | Lemma rw_ren_shift_subst_crit a u r b : 622 | eqr (r_subst (shift a u) (r_shift a r) b) 623 | (r_shift a (r_subst u r b)). 624 | Proof. 625 | repeat intro; cbn. 626 | autorewrite with rw_names. 627 | rewrite rw_ren_shift_any. 628 | rewrite rw_push_bind_shift. 629 | rewrite rw_ren_comp_id. repeat rewrite rw_apply_id. 630 | change (bind u (@applyv r (S V) (closev b v))) with (@applyv (r_subst u r b) V v). 631 | change (bind u (@applyv r (S (S V)) (closev b (wkv v)))) with (@applyv (r_subst u r b) (S V) (wkv v)). 632 | rewrite rw_applyv_wkv. 633 | names; easy. 634 | Qed. 635 | 636 | Hint Rewrite 637 | @rw_ren_id_comp 638 | @rw_ren_comp_id 639 | @rw_ren_comp_assoc 640 | @rw_ren_under_id 641 | @rw_ren_rename_rename 642 | @rw_ren_rename_rename_comp 643 | @rw_ren_rename_shift 644 | @rw_ren_rename_shift_comp 645 | @rw_ren_subst_rename 646 | @rw_ren_subst_rename_comp 647 | @rw_ren_subst_shift 648 | @rw_ren_subst_shift_comp 649 | @rw_ren_shift_any 650 | @rw_ren_shift_any_comp 651 | @rw_ren_any_subst 652 | @rw_ren_shift_subst_crit : rw_names. 653 | 654 | Goal 655 | forall a b c {V}, 656 | apply (r_shift a (r_rename b r_id c)) (vv (@free V c 0)) = 657 | vv (shiftv a (free b 0)). 658 | intros. names. 659 | (* FIXME: should shiftv be a notation? *) 660 | unfold shiftv; cbn. 661 | match goal with [ |- ?x = ?x ] => reflexivity end. 662 | Qed. 663 | 664 | Goal 665 | forall a r b {V} (t : @term V), 666 | apply (r_rename a r b) (shift b t) = shift a (apply r t). 667 | intros. names. 668 | match goal with [ |- ?x = ?x ] => reflexivity end. 669 | Qed. 670 | 671 | 672 | (* Commuting open and close with apply *) 673 | 674 | Lemma rw_push_open a r {V} (t : @term (S V)) : 675 | open a (apply r t) = apply (r_under a r) (open a t). 676 | Proof. expand_apply; names; easy. Qed. 677 | Lemma rw_push_close a r {V} (t : @term V) : 678 | apply r (close a t) = close a (apply (r_under a r) t). 679 | Proof. expand_apply; names; easy. Qed. 680 | Hint Rewrite @rw_push_open @rw_push_close : rw_names. 681 | 682 | 683 | 684 | (* 685 | * Simply-typed lambda calculus 686 | *) 687 | 688 | 689 | 690 | Inductive context {A : Set} : Set := 691 | | ctx_empty 692 | | ctx_cons (Γ : @context A) (a : string) (val : A). 693 | Notation "Γ ',,' x '~' A" := (ctx_cons Γ x A) (at level 60). 694 | 695 | Fixpoint has {A} (ctx : @context A) (x : @var 0) (v : A) := 696 | match ctx with 697 | | ctx_empty => False 698 | | ctx_cons ctx y val => 699 | match bindv (closev y x) with 700 | | None => val = v 701 | | Some x' => has ctx x' v 702 | end 703 | end. 704 | 705 | Inductive type := Base | Fn (a b : type). 706 | 707 | Reserved Notation "Γ '⊢' e '~' A" (at level 60). 708 | Inductive types (Γ : @context type) : @term 0 -> type -> Prop := 709 | | VAR x {A} : 710 | has Γ x A 711 | -> 712 | Γ ⊢ vv x ~ A 713 | | BASE k : 714 | Γ ⊢ const k ~ Base 715 | | LAM body {A B} : 716 | (forall x, (Γ,, x ~ A) ⊢ open x body ~ B) 717 | -> 718 | Γ ⊢ lam body ~ Fn A B 719 | | APP f e {A B} : 720 | Γ ⊢ f ~ Fn A B -> 721 | Γ ⊢ e ~ A -> 722 | Γ ⊢ app f e ~ B 723 | where "Γ '⊢' e '~' A" := (types Γ e A). 724 | 725 | Notation "Γ ⊆ [ r , rn ] Δ" := 726 | (forall v A, has Γ v A -> has Δ (applyrv r rn v) A) 727 | (at level 60, format "Γ ⊆ [ r , rn ] Δ"). 728 | 729 | Lemma ctx_extend_weakening {Γ Δ x} {A : type} {r} {rn : is_renaming r} : 730 | Γ ⊆ [ r , rn ] Δ -> (Γ,, x ~ A) ⊆[r_under x r,rn] (Δ,, x ~ A). 731 | Proof. 732 | intros H v B; cbn. 733 | destruct (compare x v); names; auto. 734 | Qed. 735 | 736 | Theorem weakening {Γ e A}: 737 | Γ ⊢ e ~ A -> 738 | forall r (rn : is_renaming r) Δ, 739 | Γ ⊆[r, rn] Δ -> 740 | Δ ⊢ apply r e ~ A. 741 | Proof. 742 | induction 1; cbn; intros. 743 | - erewrite applyrv_is_applyv. 744 | constructor. 745 | auto. 746 | - constructor. 747 | - constructor; intro v. 748 | names. 749 | eauto using ctx_extend_weakening. 750 | - econstructor; eauto. 751 | Qed. 752 | 753 | 754 | Notation "Γ ⊆ [ r ] Δ" := 755 | (forall v A, has Γ v A -> Δ ⊢ applyv r v ~ A) 756 | (at level 60, format "Γ ⊆ [ r ] Δ"). 757 | 758 | Lemma ctx_extend_substitution {Γ Δ x} {A : type} {r} : 759 | Γ ⊆[r] Δ -> (Γ,, x ~ A) ⊆[r_under x r] (Δ,, x ~ A). 760 | Proof. 761 | intros H v B; cbn. 762 | destruct (compare x v); names. 763 | - intros ->. constructor. names. auto. 764 | - intros ΓB. eapply (weakening (H _ _ ΓB) (r_shift x r_id) I). 765 | intros; names; auto. 766 | Qed. 767 | 768 | Theorem substitution {Γ e A} : 769 | Γ ⊢ e ~ A -> 770 | forall r Δ, 771 | Γ ⊆[r] Δ -> 772 | Δ ⊢ apply r e ~ A. 773 | Proof. 774 | induction 1; cbn; intros. 775 | - auto. 776 | - constructor. 777 | - constructor; intro v. 778 | names. 779 | eauto using ctx_extend_substitution. 780 | - econstructor; eauto. 781 | Qed. 782 | 783 | 784 | Corollary weakening1 {Γ e A B x} : 785 | Γ ⊢ e ~ A -> 786 | Γ,, x ~ B ⊢ shift x e ~ A. 787 | Proof. 788 | intro H. eapply (weakening H (r_shift x r_id) I). 789 | intros; names; auto. 790 | Qed. 791 | 792 | Corollary renaming1 {Γ e A B x y} : 793 | Γ,, x ~ B ⊢ e ~ A -> 794 | Γ,, y ~ B ⊢ apply (r_rename y r_id x) e ~ A. 795 | Proof. 796 | intro H. eapply (weakening H (r_rename y r_id x) I). 797 | intros v C. destruct (compare x v); names; auto. 798 | Qed. 799 | 800 | Corollary substitution1 {Γ e e' A B x} : 801 | Γ,, x ~ B ⊢ e ~ A -> 802 | Γ ⊢ e' ~ B -> 803 | Γ ⊢ apply (r_subst e' r_id x) e ~ A. 804 | Proof. 805 | intros He He'. eapply substitution; eauto. 806 | intros v C. destruct (compare x v); names. 807 | - intros <-; auto. 808 | - apply VAR. 809 | Qed. 810 | 811 | 812 | 813 | Lemma LAM' x {Γ body} {A B} : 814 | Γ,, x ~ A ⊢ open x body ~ B -> 815 | Γ ⊢ lam body ~ (Fn A B). 816 | Proof. 817 | intro H. 818 | apply LAM. 819 | intro y. 820 | replace (open y body) with (apply (r_rename y r_id x) (open x body)) by (names; easy). 821 | auto using renaming1. 822 | Qed. 823 | 824 | Definition swap_bound {V} (v : @var (S (S V))) : @var (S (S V)) := 825 | match v with 826 | | free vn vi => 827 | free vn vi 828 | | bound v => 829 | @bound (S (S V)) 830 | (match v with 831 | | v0 => vS v0 832 | | vS v0 => v0 833 | | vS (vS v) => vS (vS v) 834 | end) 835 | end. 836 | 837 | Lemma close_distinct {x an ai} (Hxy : x <> an) {V} : 838 | closev x (@free V an ai) = free an ai. 839 | Proof. unfold closev. string_eqs. Qed. 840 | 841 | (* This surprised me. *) 842 | Lemma swap_shift_shift {x y} {V} (v : @var V) : 843 | shiftv x (shiftv y v) = shiftv y (shiftv x v). 844 | Proof. 845 | destruct v as [vn vi|b]; unfold shiftv, openv; cbn; try easy. 846 | repeat string_eqs. 847 | Qed. 848 | 849 | Lemma shift_distinct {x y} (Hxy : x <> y) {V} : 850 | shiftv x (@free V y 0) = @free V y 0. 851 | Proof. unfold shiftv, openv; cbn. string_eqs. Qed. 852 | 853 | Lemma swap_shift_close {x y} (Hxy : x <> y) {V} (v : @var V) : 854 | closev x (shiftv y v) = shiftv y (closev x v). 855 | Proof. 856 | destruct (compare x v). 857 | - rewrite shift_distinct; autorewrite with rw_names. easy. auto. 858 | - rewrite swap_shift_shift. autorewrite with rw_names; easy. 859 | Qed. 860 | 861 | Lemma swap_close_close {x y} (Hxy : x <> y) {V} {v : @var V} : 862 | closev x (closev y v) = swap_bound (closev y (closev x v)). 863 | Proof. 864 | destruct (compare y v); autorewrite with rw_names; cbn. 865 | - replace (free y 0) with (shiftv x (@free V y 0)). 866 | autorewrite with rw_names. easy. 867 | unfold shiftv, openv; cbn. string_eqs. 868 | - rewrite swap_shift_close by assumption. autorewrite with rw_names. 869 | destruct x0 as [n i | b]; unfold closev; cbn; try easy. 870 | string_eqs. destruct i; easy. 871 | Qed. 872 | 873 | Lemma swap_subst_subst {x y} (Hxy : x <> y) {u v R} : 874 | eqr (r_subst u (r_subst v R x) y) (r_subst v (r_subst u R y) x). 875 | Proof. 876 | intros V p. cbn. 877 | rewrite (swap_close_close Hxy). 878 | destruct (closev y (closev x _)) as [pn pi|[|[|b]]]; cbn; 879 | autorewrite with rw_names; cbn; autorewrite with rw_names; easy. 880 | Qed. 881 | 882 | Lemma barendregt_substitution_debruijn x y (M : @term 0) (L N : @term 0) : 883 | x <> y -> 884 | apply (r_subst L r_id y) (apply (r_subst N r_id x) M) = 885 | apply (r_subst (apply (r_subst L r_id y) N) r_id x) (apply (r_subst (shift x L) r_id y) M). 886 | Proof. 887 | intros Hdistinct. 888 | autorewrite with rw_names. 889 | rewrite (swap_subst_subst Hdistinct); easy. 890 | Qed. 891 | 892 | Inductive Vclosed : @term 0 -> Set := 893 | | vc_vv v : Vclosed (vv v) 894 | | vc_const n : Vclosed (const n) 895 | | vc_lam t : (forall n, Vclosed (open n t)) -> Vclosed (lam t) 896 | | vc_app f e : Vclosed f -> Vclosed e -> Vclosed (app f e). 897 | 898 | Fixpoint Vclosedk (V : nat) : @term V -> Set := 899 | match V with 900 | | 0 => fun t => Vclosed t 901 | | S V => fun t => forall n, Vclosedk V (open n t) 902 | end. 903 | 904 | Fixpoint always_Vclosedk {V : nat} (t : term) {struct t} : Vclosedk V t := 905 | match t with 906 | | vv v => 907 | let fix go {V} : forall (v : @var V), Vclosedk V (vv v) := 908 | match V with 0 => vc_vv | S V => fun v n => go _ end in 909 | go v 910 | 911 | | const n => 912 | let fix go {V} : Vclosedk V _ := 913 | match V with 0 => vc_const n | S V => fun _ => go end in 914 | go 915 | 916 | | lam t => 917 | let fix go {V} : forall t, Vclosedk (S V) t -> Vclosedk V (lam t) := 918 | match V with 0 => vc_lam | S V => fun _ vc a => go _ (vc a) end in 919 | go _ (always_Vclosedk t) 920 | 921 | | app f e => 922 | let fix go {V} : forall f e, Vclosedk V f -> Vclosedk V e -> Vclosedk V (app f e) := 923 | match V with 0 => vc_app | S V => fun _ _ vf ve a => go _ _ (vf a) (ve a) end in 924 | go _ _ (always_Vclosedk f) (always_Vclosedk e) 925 | end. 926 | 927 | 928 | Lemma always_Vclosedk_open x : forall {V} (t : @term (S V)), 929 | always_Vclosedk t x = always_Vclosedk (open x t). 930 | Proof. 931 | induction_S; induction V; cbn; try easy. 932 | - rewrite IHt. easy. 933 | - rewrite IHt. easy. 934 | - rewrite IHt2. rewrite IHt1. easy. 935 | - rewrite IHt1. rewrite IHt2. easy. 936 | Qed. 937 | Hint Rewrite always_Vclosedk_open : rw_names. 938 | Open Scope string. 939 | 940 | Notation "'λ' v ~> t" := (lam (close v t)) (at level 60, format "'λ' v ~> t"). 941 | Notation "f & e" := (app f e) (at level 59, left associativity). 942 | Notation "! k" := (vv (free k 0)) (at level 10, format "! k"). 943 | 944 | Fixpoint cps {tm} (t : Vclosed tm) : @term 0 := 945 | match t with 946 | | vc_vv v => 947 | λ"k" ~> 948 | !"k" & shift "k" (vv v) 949 | 950 | | vc_const n => 951 | λ"k" ~> 952 | !"k" & const n 953 | 954 | | vc_lam _ t => 955 | λ"k" ~> 956 | !"k" & λ"x" ~> 957 | shift "k" (cps (t "x")) 958 | 959 | | vc_app _ _ f e => 960 | λ"k" ~> 961 | shift "k" (cps f) & λ"f" ~> 962 | shift "f" (shift "k" (cps e)) & λ"v" ~> 963 | !"f" & !"v" & !"k" 964 | end. 965 | Print cps. 966 | 967 | (* 968 | Meyer and Wand's theorem about typed CPS transforms 969 | ("Continuation semantics in typed lambda calculi", 1985) 970 | This follows Harper and Lillibridge's presentation. 971 | ("Polymorphic Type Assignment and CPS Conversion", 1992) 972 | *) 973 | 974 | Definition cont R A := 975 | Fn (Fn A R) R. 976 | 977 | Fixpoint cpsty R A := 978 | match A with 979 | | Base => Base 980 | | Fn A B => Fn (cpsty R A) (cont R (cpsty R B)) 981 | end. 982 | 983 | Fixpoint cpsctx R ctx {struct ctx} := 984 | match ctx with 985 | | ctx_empty => ctx_empty 986 | | ctx_cons G x A => ctx_cons (cpsctx R G) x (cpsty R A) 987 | end. 988 | 989 | Lemma cpsctx_has {R Γ A} : forall x, 990 | has Γ x A -> has (cpsctx R Γ) x (cpsty R A). 991 | Proof. 992 | induction Γ; cbn; auto. 993 | intro x; destruct (compare a x); names; [intros ->|]; auto. 994 | Qed. 995 | 996 | Theorem cps_type_preserving {Γ e A R} : 997 | types Γ e A -> 998 | types (cpsctx R Γ) (cps (always_Vclosedk e)) (cont R (cpsty R A)). 999 | Proof. 1000 | induction 1; names; intros. 1001 | - apply LAM; intro y; cbn. 1002 | eapply APP. 1003 | eapply VAR. cbn. names. easy. 1004 | eapply VAR. cbn. names. 1005 | apply cpsctx_has; auto. 1006 | - apply LAM; intro y; cbn. 1007 | eapply APP. 1008 | eapply VAR; cbn; names. easy. 1009 | apply BASE. 1010 | - apply (LAM' "k"); names. 1011 | eapply APP. apply VAR. names. easy. 1012 | apply (LAM' "x"); names. 1013 | eapply weakening. eauto. cbn. 1014 | intros v T'. cbn. names. 1015 | rewrite swap_shift_close by easy. 1016 | destruct (compare "x" v); names; easy. 1017 | - apply (LAM' "k"). names. 1018 | eapply APP. 1019 | eapply weakening1; eauto. 1020 | apply (LAM' "f"). names. 1021 | eapply APP. 1022 | eapply weakening. eauto. 1023 | intros v T'. names. easy. 1024 | apply (LAM' "v"). names. 1025 | eapply APP. 1026 | eapply APP. 1027 | eapply VAR. cbn. easy. 1028 | eapply VAR. cbn. easy. 1029 | eapply VAR. cbn. easy. 1030 | Unshelve. cbn. easy. cbn; easy. 1031 | Qed. 1032 | Print Assumptions cps_type_preserving. 1033 | 1034 | -------------------------------------------------------------------------------- /examples/Lambda.v: -------------------------------------------------------------------------------- 1 | Require Import String Morph Var Context Relative. 2 | 3 | Inductive tm {V : nat} := 4 | | vv : @var V -> @tm V 5 | | const : nat -> @tm V 6 | | lam : @tm (S V) -> @tm V 7 | | app : @tm V -> @tm V -> @tm V. 8 | 9 | Module LambdaTerm <: Term. 10 | 11 | Definition term := @tm. 12 | 13 | Definition unit {N} : morph (@var) N (@term) N := 14 | morph_inject (@vv). 15 | 16 | Fixpoint kleisli {N M} (f : morph (@var) N (@term) M) V t := 17 | match t with 18 | | vv v => f V v 19 | | const n => const n 20 | | lam t => 21 | lam (nset_push (kleisli f (S V) (nset_pop t))) 22 | | app t1 t2 => 23 | app (kleisli f V t1) (kleisli f V t2) 24 | end. 25 | 26 | Lemma left_identity : 27 | forall N M (f : morph (@var) N (@term) M) V t, 28 | kleisli f V (unit V t) = f V t. 29 | Proof. reflexivity. Qed. 30 | 31 | Lemma right_identity : 32 | forall N V (t : @term (N + V)), 33 | @kleisli N N unit V t = t. 34 | Proof. 35 | intros. 36 | inductT t; simplT; reflexivity. 37 | Qed. 38 | 39 | Lemma associativity : 40 | forall N M L 41 | (f : morph (@var) N (@term) M) 42 | (g : morph (@var) M (@term) L) V t, 43 | kleisli (fun V' t' => kleisli g V' (f V' t')) V t 44 | = kleisli g V (kleisli f V t). 45 | Proof. 46 | intros. 47 | inductT t; simplT; reflexivity. 48 | Qed. 49 | 50 | Lemma unit_extend : 51 | forall N V v, 52 | morph_extend (@unit N) V v = unit V v. 53 | Proof. 54 | intros. 55 | apply morph_extend_inject. 56 | Qed. 57 | 58 | Lemma kleisli_extend : 59 | forall N M (f : morph (@var) N (@term) M) V t, 60 | morph_extend (kleisli f) V t 61 | = kleisli (morph_extend f) V t. 62 | Proof. 63 | intros. 64 | inductT t; simplT; reflexivity. 65 | Qed. 66 | 67 | Lemma extensional : 68 | forall N M (f g : morph (@var) N (@term) M) V t, 69 | (forall V t, f V t = g V t) -> 70 | kleisli f V t = kleisli g V t. 71 | Proof. 72 | intros. 73 | inductT t; simplT; auto. 74 | Qed. 75 | 76 | End LambdaTerm. 77 | 78 | Module LambdaRen := Renamings(LambdaTerm). 79 | 80 | Import LambdaTerm. 81 | Import LambdaRen. 82 | 83 | (* 84 | * Simply-typed lambda calculus 85 | *) 86 | 87 | Inductive type := Base | Fn (a b : type). 88 | 89 | Reserved Notation "Γ '⊢' e '~' A" (at level 60). 90 | Inductive types (Γ : @context type) : @term 0 -> type -> Prop := 91 | | VAR x {A} : 92 | Γ ∋ x ~ A 93 | -> 94 | Γ ⊢ vv x ~ A 95 | | BASE k : 96 | Γ ⊢ const k ~ Base 97 | | LAM body {A B} : 98 | (forall x, Γ & x ~ A ⊢ open x body ~ B) 99 | -> 100 | Γ ⊢ lam body ~ Fn A B 101 | | APP f e {A B} : 102 | Γ ⊢ f ~ Fn A B -> 103 | Γ ⊢ e ~ A -> 104 | Γ ⊢ app f e ~ B 105 | where "Γ '⊢' e '~' A" := (types Γ e A). 106 | 107 | Hint Constructors types. 108 | 109 | Theorem weakening {Γ e A}: 110 | Γ ⊢ e ~ A -> 111 | forall r Δ, 112 | Γ =[r]=> Δ -> 113 | Δ ⊢ [r] e ~ A. 114 | Proof. 115 | induction 1; cbn; intros * rl; eauto. 116 | - rewrite applyt_related with (rl := rl). 117 | auto with contexts. 118 | - constructor; intro. 119 | names; auto with contexts. 120 | Qed. 121 | 122 | Lemma ctx_extend_substitution {Γ Δ x} {A : type} {r} : 123 | forall_has Γ (fun v B => Δ ⊢ [r] (vv v) ~ B) -> 124 | forall_has (Γ & x ~ A) 125 | (fun v B => Δ & x ~ A ⊢ [r,,x] (vv v) ~ B). 126 | Proof. 127 | intros H. 128 | apply forall_has_cons; names. 129 | - 130 | apply VAR; names; easy. 131 | - apply (forall_has_map H); intros; names. 132 | eauto using weakening with contexts. 133 | Qed. 134 | 135 | Theorem substitution {Γ e A} : 136 | Γ ⊢ e ~ A -> 137 | forall r Δ, 138 | forall_has Γ (fun v B => Δ ⊢ [r] (vv v) ~ B) -> 139 | Δ ⊢ [r] e ~ A. 140 | Proof. 141 | induction 1; cbn; intros; eauto. 142 | - constructor; intro v; names. 143 | eauto using ctx_extend_substitution. 144 | Qed. 145 | 146 | Corollary weakening1 Γ e A B x : 147 | Γ ⊢ e ~ A -> 148 | Γ & x ~ B ⊢ [^x] e ~ A. 149 | Proof. 150 | eauto using weakening with contexts. 151 | Qed. 152 | 153 | Corollary renaming1 Γ e A B x y : 154 | Γ & x ~ B ⊢ e ~ A -> 155 | Γ & y ~ B ⊢ [y <- x] e ~ A. 156 | Proof. 157 | eauto using weakening with contexts. 158 | Qed. 159 | 160 | Corollary renaming1' Γ e A B : 161 | forall x y, 162 | Γ & y ~ B ⊢ [y <- x] e ~ A -> 163 | Γ & x ~ B ⊢ e ~ A. 164 | Proof. 165 | intros x y. 166 | replace e with ([x <- y] [y <- x] e) at 2 167 | by (names; reflexivity). 168 | apply renaming1. 169 | Qed. 170 | 171 | Corollary substitution1 Γ e e' A B x : 172 | Γ & x ~ B ⊢ e ~ A -> 173 | Γ ⊢ e' ~ B -> 174 | Γ ⊢ [e' // x] e ~ A. 175 | Proof. 176 | intros He He'. 177 | apply (substitution He). 178 | apply forall_has_cons. 179 | - names; auto. 180 | - intros y b; names; auto. 181 | Qed. 182 | 183 | Lemma LAM' x {Γ body} {A B} : 184 | Γ & x ~ A ⊢ open x body ~ B -> 185 | Γ ⊢ lam body ~ (Fn A B). 186 | Proof. 187 | intro H. 188 | apply LAM. 189 | intro y. 190 | apply renaming1' with (y := x); names; easy. 191 | Qed. 192 | 193 | Lemma barendregt_substitution_debruijn x y (M : @term 0) (L N : @term 0) : 194 | distinct_names x y -> 195 | [L // y] [N // x] M = 196 | [([L // y] N) // x] [([^x] L) // y] M. 197 | Proof. 198 | intros. 199 | names. 200 | rewrite <- swap_subst_subst_distinct; easy. 201 | Qed. 202 | 203 | Inductive Vclosed : @term 0 -> Set := 204 | | vc_vv v : Vclosed (vv v) 205 | | vc_const n : Vclosed (const n) 206 | | vc_lam t : (forall n, Vclosed (open n t)) -> Vclosed (lam t) 207 | | vc_app f e : Vclosed f -> Vclosed e -> Vclosed (app f e). 208 | 209 | Fixpoint Vclosedk (V : nat) : @term V -> Set := 210 | match V with 211 | | 0 => fun t => Vclosed t 212 | | S V => fun t => forall n, Vclosedk V (open n t) 213 | end. 214 | 215 | Fixpoint always_Vclosedk {V : nat} (t : term) {struct t} : Vclosedk V t := 216 | match t with 217 | | vv v => 218 | let fix go {V} : forall (v : @var V), Vclosedk V (vv v) := 219 | match V with 0 => vc_vv | S V => fun v n => go _ end in 220 | go v 221 | | const n => 222 | let fix go {V} : Vclosedk V _ := 223 | match V with 0 => vc_const n | S V => fun _ => go end in 224 | go 225 | | lam t => 226 | let fix go {V} : forall t, Vclosedk (S V) t -> Vclosedk V (lam t) := 227 | match V with 0 => vc_lam | S V => fun _ vc a => go _ (vc a) end in 228 | go _ (always_Vclosedk t) 229 | | app f e => 230 | let fix go {V} : forall f e, Vclosedk V f -> Vclosedk V e -> Vclosedk V (app f e) := 231 | match V with 0 => vc_app | S V => fun _ _ vf ve a => go _ _ (vf a) (ve a) end in 232 | go _ _ (always_Vclosedk f) (always_Vclosedk e) 233 | end. 234 | 235 | Lemma always_Vclosedk_open x : forall {V} (t : @term (1 + V)), 236 | always_Vclosedk t x = always_Vclosedk (open x t). 237 | Proof. 238 | intros. 239 | inductT t; induction V0; cbn; try easy. 240 | - rewrite IHt; easy. 241 | - rewrite IHt; easy. 242 | - rewrite IHt1, IHt2; try apply heq_intro; easy. 243 | - rewrite IHt1, IHt2; try apply heq_intro; easy. 244 | Qed. 245 | Hint Rewrite always_Vclosedk_open : rw_names. 246 | Open Scope string. 247 | 248 | Notation "'λ' v ~> t" := (lam (close v t)) (at level 60, format "'λ' v ~> t"). 249 | Notation "f $ e" := (app f e) (at level 59, left associativity). 250 | Notation "! k" := (vv (free k)) (at level 10, format "! k"). 251 | 252 | Fixpoint cps' {tm : @term 0} (t : Vclosed tm) : @term 0 := 253 | match t with 254 | | vc_vv v => 255 | λ"k" ~> 256 | !"k" $ [^"k"] (vv v) 257 | 258 | | vc_const n => 259 | λ"k" ~> 260 | !"k" $ const n 261 | 262 | | vc_lam _ t => 263 | λ"k" ~> 264 | !"k" $ λ"x" ~> 265 | [^"k"] (cps' (t "x")) 266 | 267 | | vc_app _ _ f e => 268 | λ"k" ~> 269 | ([^"k"] (cps' f)) $ λ"f" ~> 270 | ([^"f"] ([^"k"] (cps' e))) $ λ"v" ~> 271 | !"f" $ !"v" $ !"k" 272 | end. 273 | 274 | Definition cps tm := cps' (@always_Vclosedk 0 tm). 275 | 276 | (* 277 | Meyer and Wand's theorem about typed CPS transforms 278 | ("Continuation semantics in typed lambda calculi", 1985) 279 | This follows Harper and Lillibridge's presentation. 280 | ("Polymorphic Type Assignment and CPS Conversion", 1992) 281 | *) 282 | 283 | Definition cont R A := 284 | Fn (Fn A R) R. 285 | 286 | Fixpoint cpsty R A := 287 | match A with 288 | | Base => Base 289 | | Fn A B => Fn (cpsty R A) (cont R (cpsty R B)) 290 | end. 291 | 292 | Fixpoint cpsctx R ctx {struct ctx} := 293 | match ctx with 294 | | ctx_empty => ctx_empty 295 | | ctx_cons G x A => ctx_cons (cpsctx R G) x (cpsty R A) 296 | end. 297 | 298 | Lemma cpsctx_has {R Γ} : 299 | forall_has Γ (fun x A => cpsctx R Γ ∋ x ~ cpsty R A). 300 | Proof. 301 | induction Γ; cbn; auto. 302 | apply forall_has_cons. 303 | - names; easy. 304 | - intros y b; names; auto. 305 | Qed. 306 | 307 | Theorem cps_type_preserving {Γ e A R} : 308 | Γ ⊢ e ~ A -> 309 | cpsctx R Γ ⊢ cps e ~ cont R (cpsty R A). 310 | Proof. 311 | induction 1; names; intros; apply (LAM' "k"); cbn. 312 | - eapply APP; auto with contexts. 313 | apply VAR; names. 314 | apply cpsctx_has; auto. 315 | - eauto with contexts. 316 | - eapply APP; auto with contexts. 317 | apply (LAM' "x"); names. 318 | eapply weakening; eauto. 319 | apply relates_swap_right; try easy. 320 | auto with contexts. 321 | - eapply APP; names; eauto using weakening with contexts. 322 | apply (LAM' "f"); names. 323 | eapply APP; eauto using weakening with contexts. 324 | apply (LAM' "v"); names. 325 | eauto 6 with contexts. 326 | Qed. 327 | -------------------------------------------------------------------------------- /examples/Makefile: -------------------------------------------------------------------------------- 1 | COQC=coqc 2 | COQPARAMS=-R ../src SNames 3 | 4 | MODULES=Lambda.vo 5 | 6 | all: ${MODULES} 7 | 8 | %.vo: %.v 9 | ${COQC} ${COQPARAMS} $< 10 | 11 | -------------------------------------------------------------------------------- /examples/_CoqProject: -------------------------------------------------------------------------------- 1 | -R ../src SNames 2 | -------------------------------------------------------------------------------- /src/Context.v: -------------------------------------------------------------------------------- 1 | Require Import String Morph Var. 2 | 3 | Inductive context {A : Set} : Set := 4 | | ctx_empty 5 | | ctx_cons (Γ : @context A) (a : name) (val : A). 6 | 7 | Notation "Γ & x ~ A" := (ctx_cons Γ x A) (at level 50). 8 | 9 | Fixpoint has {A} (ctx : @context A) (x : @var 0) (a : A) := 10 | match ctx with 11 | | ctx_empty => False 12 | | ctx_cons ctx y b => 13 | match bindv (closev y x) with 14 | | None => b = a 15 | | Some x' => has ctx x' a 16 | end 17 | end. 18 | 19 | Notation "Γ ∋ x ~ a" := (has Γ x a) (at level 60). 20 | 21 | Hint Extern 2 (?Γ ∋ ?x ~ ?a) => 22 | solve [cbn; easy] : contexts. 23 | 24 | (* TODO this should be done with a renaming now *) 25 | Lemma has_swap {A : Set} {Γ : @context A} {x a y b} : 26 | distinct_names x y -> 27 | forall z c, 28 | ((Γ & y ~ b) & x ~ a) ∋ z ~ c -> 29 | ((Γ & x ~ a) & y ~ b) ∋ z ~ c. 30 | Proof. 31 | intros H z c; cbn. 32 | destruct (compare_vars x z) as [|z']; 33 | autorewrite with rw_vars; cbn. 34 | - rewrite (substv_distinct (distinct_names_symmetric H)). 35 | autorewrite with rw_vars; easy. 36 | - destruct (compare_vars y z'). 37 | + autorewrite with rw_vars; cbn. 38 | rewrite rw_shift_name_distinct by easy. 39 | autorewrite with rw_vars; easy. 40 | + autorewrite with rw_vars. 41 | rewrite swap_shiftv_shiftv_distinct by easy. 42 | autorewrite with rw_vars; easy. 43 | Qed. 44 | 45 | Definition forall_has {A : Set} Γ (P : var -> A -> Prop) := 46 | (forall x a, Γ ∋ x ~ a -> P x a). 47 | 48 | Hint Unfold forall_has. 49 | 50 | Lemma forall_has_cons {A : Set} {Γ : @context A} 51 | {P : var -> A -> Prop} {x a} : 52 | P (free x) a -> 53 | forall_has Γ (fun y b => P (shiftv x y) b) -> 54 | forall_has (Γ & x ~ a) P. 55 | Proof. 56 | intros Hp Hf y b. 57 | destruct (compare_vars x y); 58 | cbn; autorewrite with rw_vars; cbn; 59 | intros; subst; auto. 60 | Qed. 61 | 62 | Lemma forall_has_map {A : Set} {Γ : @context A} 63 | {P Q : var -> A -> Prop} : 64 | forall_has Γ P -> 65 | (forall x a, P x a -> Q x a) -> 66 | forall_has Γ Q. 67 | Proof. 68 | unfold forall_has. 69 | auto. 70 | Qed. 71 | 72 | Inductive relates {trm A : Set} 73 | : @context A -> @renaming trm -> @context A -> Prop := 74 | | relates_intro : 75 | forall (Γ Δ : @context A) (s : @renaming trm) 76 | (total : total s), 77 | forall_has Γ (fun v a => has Δ (applyt s total v) a) -> 78 | relates Γ s Δ. 79 | 80 | Definition relates_total {trm A : Set} 81 | {Γ Δ : @context A} {r : @renaming trm} 82 | (rl : relates Γ r Δ) := 83 | match rl with 84 | | relates_intro _ _ _ rn _ => rn 85 | end. 86 | 87 | Definition relates_contexts {trm A : Set} 88 | {Γ Δ : @context A} {r : @renaming trm} 89 | (rl : relates Γ r Δ) := 90 | match rl in relates Γ r' Δ return 91 | (forall v a, has Γ v a -> 92 | has Δ (applyt r' (relates_total rl) v) a) with 93 | | relates_intro _ _ _ _ cs => cs 94 | end. 95 | 96 | Hint Resolve relates_total relates_contexts : contexts. 97 | 98 | Notation "Γ =[ r ]=> Δ" := 99 | (relates Γ (r)%ren Δ) (at level 60). 100 | 101 | Lemma ctx_shift {trm A : Set} : 102 | forall Γ Δ x (a : A) (r : @renaming trm), 103 | Γ =[r]=> Δ -> 104 | Γ =[r,, ^x]=> (Δ & x ~ a). 105 | Proof. 106 | intros * H. 107 | destruct H as [? ? r rn H]. 108 | apply relates_intro with (s := (r,, ^x)%ren) (total := rn). 109 | apply (forall_has_map H); intros y b. 110 | cbn; autorewrite with rw_vars rw_names. 111 | auto. 112 | Qed. 113 | 114 | Lemma ctx_rename {trm A : Set} : 115 | forall Γ Δ x y (a : A) (r : @renaming trm), 116 | Γ =[r]=> Δ -> 117 | (Γ & x ~ a) =[r,, y <- x]=> (Δ & y ~ a). 118 | Proof. 119 | intros * H. 120 | destruct H as [? ? r rn H]. 121 | apply relates_intro 122 | with (s := (r,, y <- x)%ren) (total := rn); cbn. 123 | apply forall_has_cons; 124 | autorewrite with rw_vars rw_names; cbn; auto. 125 | apply (forall_has_map H); intros z c. 126 | autorewrite with rw_vars rw_names; auto. 127 | Qed. 128 | 129 | Lemma ctx_id {trm A : Set} : 130 | forall (Γ : @context A), Γ =[(r_id : @renaming trm)]=> Γ. 131 | Proof. 132 | intros. 133 | apply relates_intro with (total := I). 134 | easy. 135 | Qed. 136 | 137 | Hint Resolve ctx_shift ctx_rename ctx_id : contexts. 138 | 139 | Lemma relates_swap_right {trm A : Set} {Γ Δ : @context A} 140 | {x a y b} {r : @renaming trm}: 141 | distinct_names x y -> 142 | Γ =[r]=> Δ & y ~ b & x ~ a -> 143 | Γ =[r]=> Δ & x ~ a & y ~ b. 144 | Proof. 145 | intros Heq H. 146 | inversion H as [? ? ? tl H']; subst. 147 | apply relates_intro with (total := tl). 148 | auto using has_swap. 149 | Qed. 150 | 151 | Lemma relates_swap_left {trm A : Set} {Γ Δ : @context A} 152 | {x a y b} {r : @renaming trm}: 153 | distinct_names y x -> 154 | Γ & y ~ b & x ~ a =[r]=> Δ -> 155 | Γ & x ~ a & y ~ b =[r]=> Δ. 156 | Proof. 157 | intros Heq H. 158 | inversion H as [? ? ? tl H']; subst. 159 | apply relates_intro with (total := tl). 160 | auto using has_swap. 161 | Qed. 162 | -------------------------------------------------------------------------------- /src/Makefile: -------------------------------------------------------------------------------- 1 | COQC=coqc 2 | COQPARAMS=-R . SNames 3 | 4 | MODULES=Morph.vo Var.vo Context.vo Relative.vo 5 | 6 | all: ${MODULES} 7 | 8 | %.vo: %.v 9 | ${COQC} ${COQPARAMS} $< 10 | 11 | Var.vo: Morph.vo 12 | 13 | Context.vo: Var.vo 14 | 15 | Relative.vo: Morph.vo Var.vo Context.vo 16 | 17 | -------------------------------------------------------------------------------- /src/Morph.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.EqdepFacts. 2 | Require Import Coq.Logic.Eqdep_dec. 3 | Require Import Coq.Arith.Peano_dec. 4 | 5 | Definition nset := forall (V : nat), Set. 6 | 7 | Definition heq {T : nset} : 8 | forall {N : nat}, @T N -> forall {M : nat}, @T M -> Prop := 9 | eq_dep nat (@T). 10 | 11 | Definition heq_intro := eq_dep_intro. 12 | 13 | Hint Resolve heq_intro. 14 | 15 | Notation " x ~= y " := 16 | (heq x y) (at level 70, no associativity). 17 | 18 | Definition cast {T : nset} {N M} (pf : N = M) (t : @T N) : @T M := 19 | match pf in (_ = L) return (@T L) with 20 | | eq_refl => t 21 | end. 22 | 23 | Lemma eq_heq : forall {T : nset} {N} {t s : @T N}, 24 | t = s -> t ~= s. 25 | Proof. 26 | intros T N t s H. 27 | rewrite H. 28 | apply eq_dep_intro. 29 | Qed. 30 | 31 | Definition heq_eq : 32 | forall {T : nset} {N} {t s : @T N}, t ~= s -> t = s := 33 | eq_dep_eq_dec eq_nat_dec. 34 | 35 | Lemma heq_trans : 36 | forall (T:nset) N M L (t : @T N) (s : @T M) (r : @T L), 37 | t ~= s -> s ~= r -> t ~= r. 38 | Proof. 39 | unfold heq. 40 | eauto using eq_dep_trans. 41 | Qed. 42 | 43 | Lemma heq_under : forall T N M t1 t2, 44 | @heq T (S N) t1 (S M) t2 45 | <-> @heq (fun V : nat => T (S V)) N t1 M t2. 46 | Proof. 47 | unfold heq. 48 | intros T N M t1 t2. 49 | split; intro H; 50 | inversion H; subst; 51 | apply heq_eq in H; subst; 52 | apply eq_dep_intro. 53 | Qed. 54 | 55 | Definition push_eq N V := 56 | nat_ind (fun N' : nat => N' + S V = S (N' + V)) 57 | (@eq_refl nat (S V)) 58 | (fun (N' : nat) (IHn : N' + S V = S (N' + V)) => 59 | f_equal_nat nat S (N' + S V) (S (N' + V)) IHn) N. 60 | 61 | Definition pop_eq N V := eq_sym (push_eq N V). 62 | 63 | Definition nset_push {T : nset} {N V} 64 | (t : @T (N + S V)) : @T (S (N + V)) := 65 | cast (push_eq N V) t. 66 | 67 | Definition nset_pop {T : nset} {N V} 68 | (t : @T (S (N + V))) : @T (N + S V) := 69 | cast (pop_eq N V) t. 70 | 71 | Lemma nset_push_heq : 72 | forall (T : nset) N V (t : @T (N + S V)), 73 | nset_push t ~= t. 74 | Proof. 75 | intros. 76 | unfold nset_push, cast. 77 | destruct (push_eq N V). 78 | apply eq_dep_intro. 79 | Qed. 80 | 81 | Lemma nset_pop_heq : 82 | forall (T : nset) N V (t : @T (S (N + V))), 83 | nset_pop t ~= t. 84 | Proof. 85 | intros. 86 | unfold nset_pop, cast. 87 | destruct (pop_eq N V). 88 | apply eq_dep_intro. 89 | Qed. 90 | 91 | Lemma nset_push_pop_eq : 92 | forall (T : nset) N V (t : @T (S (N + V))), 93 | nset_push (nset_pop t) = t. 94 | Proof. 95 | intros T N V t. 96 | unfold nset_push, nset_pop, cast, pop_eq. 97 | destruct (push_eq N V); cbn. 98 | reflexivity. 99 | Qed. 100 | 101 | Lemma nset_pop_push_eq : 102 | forall (T : nset) N V (t : @T (N + S V)), 103 | nset_pop (nset_push t) = t. 104 | Proof. 105 | intros T N V t. 106 | unfold nset_push, nset_pop, cast, pop_eq. 107 | destruct (push_eq N V); cbn. 108 | reflexivity. 109 | Qed. 110 | 111 | Lemma nset_pop_under : forall T N V t, 112 | @nset_pop (fun N' : nat => T (S N')) N V t 113 | = @nset_pop T (S N) V t. 114 | Proof. 115 | intros. 116 | apply heq_eq. 117 | apply heq_trans with (s := t). 118 | - rewrite heq_under. 119 | apply nset_pop_heq. 120 | - apply eq_dep_sym. 121 | apply nset_pop_heq with (N := S N). 122 | Qed. 123 | 124 | Lemma nset_push_under : forall T N V t, 125 | @nset_push (fun N' : nat => T (S N')) N V t 126 | = @nset_push T (S N) V t. 127 | Proof. 128 | intros. 129 | apply heq_eq. 130 | apply heq_trans with (s := t). 131 | - rewrite heq_under. 132 | apply nset_push_heq. 133 | - apply eq_dep_sym. 134 | apply nset_push_heq with (N := S N). 135 | Qed. 136 | 137 | Definition morph (T : nset) (N : nat) (S : nset) (M : nat) := 138 | forall V, @T (N + V) -> @S (M + V). 139 | 140 | Definition morph_inject {T S: nset} {N} 141 | (f : forall V, @T V -> @S V) 142 | : morph (@T) N (@S) N := fun V t => f (N + V) t. 143 | 144 | Arguments morph_inject {T S N} f /. 145 | 146 | Definition morph_id {T N} : morph (@T) N (@T) N := 147 | (fun _ t => t). 148 | 149 | Arguments morph_id {T N} V t /. 150 | 151 | Notation " 1 " := morph_id : morph_scope. 152 | 153 | Definition morph_compose {T N S M R L} : 154 | morph (@S) M (@R) L -> 155 | morph (@T) N (@S) M -> 156 | morph (@T) N (@R) L := 157 | fun m2 m1 => 158 | fun V t => m2 V (m1 V t). 159 | 160 | Arguments morph_compose {T N S M R L} m1 m2 V t /. 161 | 162 | Notation "m1 @ m2" := (morph_compose m1 m2) 163 | (at level 60, right associativity) 164 | : morph_scope. 165 | 166 | Bind Scope morph_scope with morph. 167 | Open Scope morph_scope. 168 | 169 | Lemma morph_left_identity : 170 | forall T N S M (f : morph (@T) N (@S) M), 171 | 1 @ f = f. 172 | Proof. reflexivity. Qed. 173 | 174 | Lemma morph_right_identity : 175 | forall T N S M (f : morph (@T) N (@S) M), 176 | f @ 1 = f. 177 | Proof. reflexivity. Qed. 178 | 179 | Lemma morph_associative : 180 | forall T N S M R L U O 181 | (f : morph (@T) N (@S) M) 182 | (g : morph (@R) L (@T) N) 183 | (h : morph (@U) O (@R) L), 184 | f @ (g @ h) = (f @ g) @ h. 185 | Proof. reflexivity. Qed. 186 | 187 | (* Extension *) 188 | 189 | Definition morph_extend {T N R L} (m : morph (@T) N (@R) L) 190 | : morph (@T) (S N) (@R) (S L) := 191 | fun V t => nset_push (m (S V) (nset_pop t)). 192 | 193 | (* Automation *) 194 | 195 | Ltac inductT t := 196 | match type of t with 197 | | context T [?N + ?V] => 198 | let t' := fresh "t" in 199 | let NV := fresh "NV" in 200 | let Heq := fresh "Heq" in 201 | let HeqNV := fresh "HeqNV" in 202 | let V' := fresh "V" in 203 | remember t as t' eqn:Heq; 204 | apply eq_heq in Heq; 205 | remember (N + V) as NV eqn:HeqNV in t, Heq at 2; 206 | generalize dependent HeqNV; 207 | generalize dependent t'; 208 | generalize dependent V; 209 | induction t; intros V' t' Heq HeqNV; 210 | subst; rewrite (heq_eq Heq); clear Heq; cbn 211 | | context T [?N + ?V] => 212 | fail "unexpected failure" 213 | | _ => 214 | fail "term's type is not of the form '@T (?N + ?V)'" 215 | end. 216 | 217 | Ltac pop_term_arguments t := 218 | match t with 219 | | ?f ?s => 220 | match type of s with 221 | | context T [S (?N + ?V)] => 222 | let R := 223 | constr:(fun N => 224 | ltac:(let y := context T[N] in exact y)) 225 | in 226 | assert (@nset_pop R N V s ~= s) 227 | by apply nset_pop_heq; 228 | generalize dependent (@nset_pop R N V s); 229 | pop_term_arguments f; 230 | match goal with 231 | | Heq : _ ~= s |- _ => 232 | rewrite (heq_eq Heq); clear Heq 233 | end 234 | | nat => 235 | match s with 236 | | S (?N + ?V) => 237 | replace (N + S V) with s 238 | by apply (pop_eq N V); 239 | intros 240 | | _ => pop_term_arguments f 241 | end 242 | | _ => pop_term_arguments f 243 | end 244 | end. 245 | 246 | Ltac popped_term t := 247 | match t with 248 | | ?f ?s => 249 | let f' := popped_term f in 250 | match type of s with 251 | | context T [S (?N + ?V)] => 252 | let R := 253 | constr:(fun N => 254 | ltac:(let y := context T[N] in exact y)) 255 | in 256 | constr:(f' (@nset_pop R N V s)) 257 | | nat => 258 | match s with 259 | | S (?N + ?V) => constr:(f' (N + S V)) 260 | | _ => constr:(f' s) 261 | end 262 | | _ => 263 | constr:(f' s) 264 | end 265 | | ?f => 266 | constr:(f) 267 | end. 268 | 269 | Ltac popT := 270 | cbn in *; 271 | match goal with 272 | | |- context T [@nset_pop ?T' ?N ?V ?t] => 273 | let t' := popped_term t in 274 | replace (@nset_pop T' N V t) with t'; 275 | [> | symmetry; apply heq_eq; 276 | apply eq_dep_trans with (y := t); 277 | [>apply nset_pop_heq|]; 278 | pop_term_arguments t; 279 | reflexivity 280 | ] 281 | end; 282 | repeat rewrite nset_pop_under. 283 | 284 | Ltac push_term_arguments t := 285 | match t with 286 | | ?f ?s => 287 | match type of s with 288 | | context T [?N + S ?V] => 289 | let R := 290 | constr:(fun N => 291 | ltac:(let y := context T[N] in exact y)) 292 | in 293 | assert (@nset_push R N V s ~= s) 294 | by apply nset_push_heq; 295 | generalize dependent (@nset_push R N V s); 296 | push_term_arguments f; 297 | match goal with 298 | | Heq : _ ~= s |- _ => 299 | rewrite (heq_eq Heq); clear Heq 300 | end 301 | | nat => 302 | match s with 303 | | (?N + S ?V) => 304 | replace (S (N + V)) with s 305 | by apply (push_eq N V); 306 | intros 307 | | _ => push_term_arguments f 308 | end 309 | | _ => push_term_arguments f 310 | end 311 | end. 312 | 313 | Ltac pushed_term t := 314 | match t with 315 | | ?f ?s => 316 | let f' := pushed_term f in 317 | match type of s with 318 | | context T [?N + S ?V] => 319 | let R := 320 | constr:(fun N => 321 | ltac:(let y := context T[N] in exact y)) 322 | in 323 | constr:(f' (@nset_push R N V s)) 324 | | nat => 325 | match s with 326 | | (?N + S ?V) => constr:(f' (S (N + V))) 327 | | _ => constr:(f' s) 328 | end 329 | | _ => 330 | constr:(f' s) 331 | end 332 | | ?f => 333 | constr:(f) 334 | end. 335 | 336 | Ltac pushT := 337 | cbn in *; 338 | match goal with 339 | | |- context T [@nset_push ?T' ?N ?V ?t] => 340 | let t' := pushed_term t in 341 | replace (@nset_push T' N V t) with t'; 342 | [> | symmetry; apply heq_eq; 343 | apply eq_dep_trans with (y := t); 344 | [>apply nset_push_heq|]; 345 | push_term_arguments t; 346 | reflexivity 347 | ] 348 | end; 349 | repeat rewrite nset_push_under. 350 | 351 | Ltac simplT := 352 | unfold morph_extend in *; 353 | try popT; 354 | try pushT; 355 | repeat 356 | match goal with 357 | | IH : forall (_ : nat) (_ : _), 358 | _ ~= _ -> _ = _ -> _ = _ |- _ => 359 | rewrite IH; 360 | [> | 361 | match goal with 362 | | |- @nset_pop ?T' ?N' ?V' ?t' ~= ?t' => 363 | apply nset_pop_heq with (N := N') 364 | | |- ?t' ~= ?t' => 365 | apply heq_intro 366 | end 367 | | auto ] 368 | end; 369 | try rewrite nset_push_pop_eq; 370 | try rewrite nset_pop_push_eq. 371 | 372 | Lemma morph_extend_inject : 373 | forall (T R : nset) N 374 | (f : forall V : nat, T V -> R V) V t, 375 | @morph_extend T N R N (morph_inject f) V t 376 | = morph_inject f V t. 377 | Proof. 378 | intros. 379 | unfold morph_extend, morph_inject. 380 | pushT. 381 | rewrite nset_push_pop_eq. 382 | reflexivity. 383 | Qed. 384 | -------------------------------------------------------------------------------- /src/Relative.v: -------------------------------------------------------------------------------- 1 | Require Import String Morph Var Context. 2 | Require Setoid Morphisms. 3 | 4 | Module Type Term. 5 | 6 | Parameter term : forall {V : nat}, Set. 7 | 8 | Parameter unit : forall {N}, morph (@var) N (@term) N. 9 | 10 | Parameter kleisli : 11 | forall {N M}, 12 | morph (@var) N (@term) M -> 13 | morph (@term) N (@term) M. 14 | 15 | Axiom left_identity : 16 | forall N M (f : morph (@var) N (@term) M) V t, 17 | kleisli f V (unit V t) = f V t. 18 | 19 | Axiom right_identity : 20 | forall N V (t : @term (N + V)), 21 | kleisli unit V t = t. 22 | 23 | Axiom associativity : 24 | forall N M L 25 | (f : morph (@var) N (@term) M) 26 | (g : morph (@var) M (@term) L) V t, 27 | kleisli (fun V' t' => kleisli g V' (f V' t')) V t 28 | = kleisli g V (kleisli f V t). 29 | 30 | Axiom unit_extend : 31 | forall N V v, 32 | morph_extend (@unit N) V v = unit V v. 33 | 34 | Axiom kleisli_extend : 35 | forall N M (f : morph (@var) N (@term) M) V t, 36 | morph_extend (kleisli f) V t 37 | = kleisli (morph_extend f) V t. 38 | 39 | Axiom extensional : 40 | forall N M (f g : morph (@var) N (@term) M) V t, 41 | (forall V t, f V t = g V t) -> 42 | kleisli f V t = kleisli g V t. 43 | 44 | End Term. 45 | 46 | Module Renamings (T : Term). 47 | 48 | Import T. 49 | 50 | Definition openm a := 51 | unit @ ((@openv a : morph (@var) 1 (@var) 0)). 52 | 53 | Definition open a := kleisli (openm a). 54 | 55 | Definition closem a := 56 | unit @ ((@closev a : morph (@var) 0 (@var) 1)). 57 | 58 | Definition close a := kleisli (closem a). 59 | 60 | Definition wkm := 61 | unit @ (@wkv : morph (@var) 0 (@var) 1). 62 | 63 | Definition wk := kleisli wkm. 64 | 65 | Fixpoint wkn {V} (t : @term 0) := 66 | match V with 67 | | 0 => t 68 | | S V => wk V (@wkn V t) 69 | end. 70 | 71 | Definition bindm (u : @term 0) : morph (@var) 1 (@term) 0 := 72 | fun V v => 73 | match bindv v with 74 | | Some v => @unit 0 V v 75 | | None => wkn u 76 | end. 77 | 78 | Definition bind u := kleisli (@bindm u). 79 | 80 | Arguments openm a {V} v /. 81 | Arguments open a {V} !t. 82 | Arguments closem a {V} v /. 83 | Arguments close a {V} !t. 84 | Arguments wkm {V} v /. 85 | Arguments wk {V} !t. 86 | Arguments bindm u {V} !v. 87 | Arguments bind u {V} !t. 88 | 89 | Definition ren := @renaming (@term 0). 90 | 91 | Fixpoint applyv (r : ren) : morph (@var) 0 (@term) 0 := 92 | match r with 93 | | r_id => unit 94 | | r_comp r s => 95 | (kleisli (applyv r)) 96 | @ (applyv s) 97 | | r_shift b r => 98 | (@open b) 99 | @ (kleisli (morph_extend (applyv r))) 100 | @ (@wkm) 101 | | r_rename b r a => 102 | (@open b) 103 | @ (kleisli (morph_extend (applyv r))) 104 | @ (@closem a) 105 | | r_subst u r a => 106 | (@bind u) 107 | @ (kleisli (morph_extend (applyv r))) 108 | @ (@closem a) 109 | end. 110 | 111 | Arguments applyv !r {V} t /. 112 | 113 | Notation "[ r ] t" := (kleisli (@applyv (r)%ren) _ t) 114 | (at level 60, right associativity) : term_scope. 115 | Bind Scope term_scope with term. 116 | Open Scope term_scope. 117 | 118 | Lemma unit_extend_0 : forall V t, 119 | @unit 0 (S V) t = @unit 1 V t. 120 | Proof. 121 | intros. 122 | replace (@unit 0 (S V) t) 123 | with (morph_extend (@unit 0) V t) by reflexivity. 124 | rewrite unit_extend. 125 | reflexivity. 126 | Qed. 127 | 128 | Lemma kleisli_extend_00 : forall V f t, 129 | @kleisli 0 0 f (S V) t 130 | = @kleisli 1 1 (morph_extend f) V t. 131 | Proof. 132 | intros. 133 | replace (@kleisli 0 0 f (S V) t) 134 | with (morph_extend (@kleisli 0 0 f) V t) 135 | by reflexivity. 136 | rewrite kleisli_extend. 137 | reflexivity. 138 | Qed. 139 | 140 | Lemma right_identity_ext : 141 | forall N V (t : @term (N + V)) f, 142 | (forall V' v', f V' v' = unit V' v') -> 143 | kleisli f V t = t. 144 | Proof. 145 | intros * Heq. 146 | rewrite extensional with (g := unit); 147 | auto using right_identity. 148 | Qed. 149 | 150 | Ltac simpl_term t := 151 | match t with 152 | | @kleisli ?N _ _ ?V (@kleisli _ ?N _ ?V _) => 153 | rewrite <- associativity 154 | | @kleisli 0 0 _ (S _) _ => 155 | rewrite kleisli_extend_00 156 | | @kleisli ?N _ _ ?V (@unit ?N ?V _) => 157 | rewrite left_identity 158 | | @kleisli ?N ?N (@unit ?N) ?V _ => 159 | rewrite right_identity 160 | | @unit 0 (S _) _ => 161 | rewrite unit_extend_0 162 | | morph_compose _ _ _ _ => unfold morph_compose 163 | | morph_compose _ _ => unfold morph_compose; intros 164 | | @open _ _ _ => unfold open 165 | | @openm _ _ _ => unfold openm 166 | | @close _ _ _ => unfold close 167 | | @closem _ _ _ => unfold closem 168 | | @wk _ _ => unfold wk 169 | | @wkm _ _ => unfold wkm 170 | | @bind _ _ _ => unfold bind 171 | | @bindm _ _ _ => unfold bindm 172 | | kleisli _ _ ?t' => 173 | simpl_term t' 174 | end. 175 | 176 | Ltac simpl_term_eq_step := 177 | match goal with 178 | | |- ?t1 = _ => simpl_term t1; cbn 179 | | |- _ = ?t2 => simpl_term t2; cbn 180 | | |- @kleisli ?N ?M _ ?V ?t = @kleisli ?N ?M _ ?V ?t => 181 | apply (extensional N M); intros; cbn 182 | | |- @kleisli ?N ?N _ _ ?t = ?t => 183 | apply (right_identity_ext N); intros; cbn 184 | | |- ?t = @kleisli ?N ?N _ _ ?t => 185 | symmetry; apply (right_identity_ext N); intros; 186 | symmetry; cbn 187 | | |- _ = _ => 188 | progress autorewrite with rw_names rw_vars; cbn 189 | end. 190 | 191 | Ltac simpl_term_eq := 192 | cbn; repeat simpl_term_eq_step; try reflexivity. 193 | 194 | Lemma rw_close_open a : 195 | forall {V} (t : @term (S V)), close a (open a t) = t. 196 | Proof. 197 | intros. 198 | simpl_term_eq. 199 | Qed. 200 | 201 | Lemma rw_open_close a : 202 | forall {V} (t : @term V), open a (close a t) = t. 203 | Proof. 204 | intros. 205 | simpl_term_eq. 206 | Qed. 207 | 208 | Lemma rw_bind_wk u : 209 | forall {V} (t : @term V), bind u (wk t) = t. 210 | Proof. 211 | intros. 212 | simpl_term_eq. 213 | Qed. 214 | 215 | Lemma rw_bindm_wkv u : 216 | forall {V} (v : @var V), bindm u (wkv v) = @unit 0 V v. 217 | Proof. 218 | intros. 219 | simpl_term_eq. 220 | Qed. 221 | 222 | Hint Rewrite @rw_close_open @rw_open_close @rw_bind_wk 223 | @rw_bindm_wkv 224 | : rw_renaming. 225 | 226 | (* identity and composition for apply *) 227 | Lemma rw_apply_id {V} (t : @term V) : [r_id] t = t. 228 | Proof. simpl_term_eq. Qed. 229 | 230 | Lemma rw_apply_comp r s {V} (t : @term V) : 231 | [r] [s] t = [r;s] t. 232 | Proof. simpl_term_eq. Qed. 233 | 234 | Hint Rewrite @rw_apply_id @rw_apply_comp : rw_renaming. 235 | 236 | (* Simple rewritings on applyv and shiftv (more later) *) 237 | Lemma rw_applyv_bound r : forall {V} (v : level V), 238 | @applyv r V (bound v) = @unit 0 V (bound v). 239 | Proof. 240 | intros. 241 | generalize dependent V. 242 | induction r; intros; simpl_term_eq. 243 | - rewrite IHr2; simpl_term_eq. 244 | rewrite IHr1; simpl_term_eq. 245 | - rewrite IHr; simpl_term_eq. 246 | - rewrite IHr; simpl_term_eq. 247 | - rewrite IHr; simpl_term_eq. 248 | Qed. 249 | 250 | Hint Rewrite @rw_applyv_bound : rw_renaming. 251 | 252 | (* Rewrite to do the following: 253 | - group balanced pairs of operations into 'ren's 254 | - push opens/bind right and wk/close left 255 | - simplify rens *) 256 | 257 | Lemma rw_group_rename a b {V} (t : @term V) : 258 | open a (close b t) = [a <- b] t. 259 | Proof. simpl_term_eq. Qed. 260 | 261 | Lemma rw_group_shift a {V} (t : @term V) : 262 | open a (wk t) = [^a] t. 263 | Proof. simpl_term_eq. Qed. 264 | 265 | Lemma rw_group_subst u a {V} (t : @term V) : 266 | bind u (close a t) = [u // a] t. 267 | Proof. simpl_term_eq. Qed. 268 | 269 | Hint Rewrite @rw_group_rename @rw_group_shift 270 | @rw_group_subst 271 | : rw_renaming. 272 | 273 | (* fold operations *) 274 | 275 | Lemma rw_fold_open a {V} (t : @term (1 + V)) : 276 | kleisli (@openm a) V t = open a t. 277 | Proof. easy. Qed. 278 | 279 | Lemma rw_fold_close a {V} (t : @term V) : 280 | kleisli (@closem a) V t = close a t. 281 | Proof. easy. Qed. 282 | 283 | Lemma rw_fold_wk {V} (t : @term V) : 284 | kleisli (@wkm) V t = wk t. 285 | Proof. easy. Qed. 286 | 287 | Lemma rw_fold_bind u {V} (t : @term (1 + V)) : 288 | kleisli (@bindm u) V t = bind u t. 289 | Proof. easy. Qed. 290 | 291 | Hint Rewrite @rw_fold_open @rw_fold_close @rw_fold_wk 292 | @rw_fold_bind 293 | : rw_renaming. 294 | 295 | (* wk commutes with shift *) 296 | Lemma rw_shift_wk b {V} (t : @term V) : 297 | [^b] (wk t) = wk ([^b] t). 298 | Proof. 299 | simpl_term_eq. 300 | rewrite swap_shiftv_wkv. 301 | simpl_term_eq. 302 | Qed. 303 | 304 | Hint Rewrite @rw_shift_wk : rw_renaming. 305 | 306 | (* wk commutes with apply. 307 | Somewhat harder to prove than I'd expect *) 308 | Lemma rw_applyv_wkv r : forall {V} (v : @var V), 309 | applyv r (wkv v) = wk (applyv r v). 310 | Proof. 311 | induction r; intros; cbn. 312 | - simpl_term_eq. 313 | - rewrite IHr2; simpl_term_eq. 314 | rewrite IHr1; simpl_term_eq. 315 | - simpl_term_eq. 316 | repeat rewrite IHr; simpl_term_eq. 317 | rewrite swap_shiftv_wkv. 318 | simpl_term_eq. 319 | - destruct (compare_vars a v); simpl_term_eq. 320 | + autorewrite with rw_renaming. 321 | simpl_term_eq. 322 | + repeat rewrite IHr. 323 | simpl_term_eq. 324 | rewrite swap_shiftv_wkv. 325 | simpl_term_eq. 326 | - destruct (compare_vars a v); simpl_term_eq. 327 | + autorewrite with rw_renaming. 328 | simpl_term_eq. 329 | + repeat rewrite IHr; simpl_term_eq. 330 | Qed. 331 | 332 | (* FIXME where? *) 333 | Lemma applyt_is_applyv : 334 | forall r (rn : total r) {V} (v : @var V), 335 | applyv r v = unit V (applyt r rn v). 336 | Proof. 337 | induction r; cbn; intuition; simpl_term_eq. 338 | - erewrite IHr2; simpl_term_eq. 339 | erewrite IHr1; simpl_term_eq. 340 | - erewrite IHr; simpl_term_eq. 341 | - erewrite IHr; simpl_term_eq. 342 | Qed. 343 | 344 | Lemma applyt_related (A : Set) (Γ Δ : @context A) : 345 | forall r (rl : Γ =[r]=> Δ) {V} (v : @var V), 346 | applyv r v = unit V (applyt r (relates_total rl) v). 347 | Proof. 348 | intros. 349 | rewrite applyt_is_applyv with (rn := relates_total rl). 350 | reflexivity. 351 | Qed. 352 | 353 | Lemma rw_apply_wk r {V} (t : @term V) : 354 | kleisli (@applyv r) (S V) (wk t) = wk ([r] t). 355 | Proof. 356 | simpl_term_eq. 357 | apply rw_applyv_wkv. 358 | Qed. 359 | 360 | Lemma rw_apply_wkn r {V} (t : @term 0) : 361 | [r] (@wkn V t) = @wkn V ([r] t). 362 | Proof. 363 | induction V; cbn; try easy. 364 | rewrite rw_apply_wk. 365 | rewrite IHV. 366 | reflexivity. 367 | Qed. 368 | 369 | Hint Rewrite @rw_applyv_wkv @rw_apply_wk @rw_apply_wkn 370 | : rw_renaming. 371 | 372 | Lemma rw_applyv_wkv_free r {a V} : 373 | applyv r (@free (S V) a) = wk (applyv r (@free V a)). 374 | Proof. 375 | change (@free (S V) a) with (wkv (@free V a)). 376 | apply rw_applyv_wkv. 377 | Qed. 378 | 379 | Hint Rewrite @rw_applyv_wkv_free : rw_renaming. 380 | 381 | Lemma comm_bind_apply u : forall {V} (t : @term (S V)) r, 382 | [r] (bind u t) = bind ([r] u) ([r] t). 383 | Proof. 384 | intros. 385 | simpl_term_eq. 386 | destruct (compare_l0 t0); simpl_term_eq; 387 | autorewrite with rw_renaming; simpl_term_eq. 388 | Qed. 389 | 390 | (* FIXME: maybe this should just be "names"? *) 391 | Lemma rw_close_shift a r {V} (t : @term V) : 392 | close a ([r,, ^a] t) = wk ([r] t). 393 | Proof. 394 | simpl_term_eq; autorewrite with rw_renaming; simpl_term_eq. 395 | Qed. 396 | Lemma rw_subst_open u r b {V} (t : @term (S V)) : 397 | [r,, u // b] (open b t) = bind u ([r] t). 398 | Proof. simpl_term_eq. Qed. 399 | Lemma rw_rename_open a b {V} (t : @term (S V)) : 400 | [a <- b] (open b t) = open a t. 401 | Proof. simpl_term_eq. Qed. 402 | Hint Rewrite @rw_close_shift @rw_subst_open @rw_rename_open : rw_renaming. 403 | 404 | (* FIXME: lots of applyt rewritings below here *) 405 | 406 | (* FIXME: not great, tbh *) 407 | 408 | (* Commuting bind with apply. 409 | It seems more natural to make 410 | bind (r u) . r --> r . bind u 411 | since bind is similar to open, and open moves right. 412 | But that rule breaks confluence, and isn't very useful 413 | as bind doesn't show up in terms. So, binds move left 414 | instead. *) 415 | Lemma rw_push_bind u r {V} (t : @term (S V)) : 416 | [r] (bind u t) = bind ([r] u) ([r] t). 417 | Proof. apply comm_bind_apply. Qed. 418 | (* Shifts go the other way. 419 | This is confluent/terminating because shift doesn't simplify to apply (?) *) 420 | Lemma rw_push_bind_shift u r s a {V} (t : @term (S V)) : 421 | bind ([r,, ^a] u) ([s,, ^ a] t) = 422 | [^a] (bind ([r]u) ([s]t)). 423 | Proof. 424 | rewrite comm_bind_apply. 425 | autorewrite with rw_renaming; simpl_term_eq. 426 | autorewrite with rw_renaming; simpl_term_eq. 427 | destruct (compare_l0 t1); simpl_term_eq. 428 | f_equal. 429 | simpl_term_eq. 430 | autorewrite with rw_renaming; simpl_term_eq. 431 | Qed. 432 | Lemma rw_push_bind_shiftvu v a {V} (t : @term (S V)) : 433 | bind (@unit 0 0 (shiftv a v)) ([^a] t) 434 | = [^a] (bind (unit 0 v) t). 435 | Proof. 436 | replace (unit 0 (shiftv a v)) with ([^a] (unit 0 v)) 437 | by simpl_term_eq. 438 | rewrite rw_push_bind_shift. 439 | autorewrite with rw_renaming; easy. 440 | Qed. 441 | 442 | Hint Rewrite 443 | @rw_push_bind @rw_push_bind_shift @rw_push_bind_shiftvu 444 | : rw_renaming. 445 | 446 | (* Morphisms. *) 447 | Import Setoid Morphisms. 448 | 449 | (* FIXME: is it better to define eqr over all vars or just free ones? *) 450 | Definition eqr (r s : ren) : Prop := 451 | forall V (v : @var V), 452 | applyv r v = applyv s v. 453 | 454 | Instance eqr_Equivalence : Equivalence eqr. 455 | Proof. 456 | split; red; unfold eqr; auto. 457 | intros. etransitivity; eauto. 458 | Qed. 459 | 460 | Instance mor_applyv : Proper 461 | (eqr ==> forall_relation (fun _ => (eq ==> eq))) applyv. 462 | Proof. intros r s Hrs V v v' <-. auto. Qed. 463 | 464 | Instance mor_kleisli : Proper 465 | (forall_relation (fun _ => (eq ==> eq)) 466 | ==> forall_relation (fun _ => (eq ==> eq))) 467 | (@kleisli 0 0). 468 | Proof. 469 | intros r s Hrs V a b <-. 470 | simpl_term_eq. 471 | apply Hrs; reflexivity. 472 | Qed. 473 | 474 | Ltac easy_eqr := 475 | repeat intro; simpl_term_eq; 476 | repeat 477 | match goal with 478 | | [ H : _ |- _ ] => 479 | rewrite H; clear H; simpl_term_eq 480 | end. 481 | 482 | Instance mor_r_comp : Proper 483 | (eqr ==> eqr ==> eqr) r_comp. easy_eqr. Qed. 484 | Instance mor_r_shift : Proper 485 | (eq ==> eqr ==> eqr) r_shift. easy_eqr. Qed. 486 | Instance mor_r_rename : Proper 487 | (eq ==> eqr ==> eq ==> eqr) r_rename. easy_eqr. Qed. 488 | Instance mor_r_subst : Proper 489 | (eq ==> eqr ==> eq ==> eqr) r_subst. easy_eqr. Qed. 490 | 491 | Lemma rw_ren_id_comp r : eqr (r_comp r_id r) r. 492 | Proof. easy_eqr. Qed. 493 | Lemma rw_ren_comp_id r : eqr (r_comp r r_id) r. 494 | Proof. easy_eqr. Qed. 495 | Lemma rw_ren_comp_assoc r s t : 496 | eqr (r_comp r (r_comp s t)) (r_comp (r_comp r s) t). 497 | Proof. easy_eqr. Qed. 498 | Lemma rw_ren_under_id x : 499 | eqr (r_rename x r_id x) r_id. 500 | Proof. easy_eqr. Qed. 501 | Lemma rw_ren_rename_rename c b a r s : 502 | eqr (r_comp (r_rename c r b) (r_rename b s a)) 503 | (r_rename c (r_comp r s) a). 504 | Proof. easy_eqr. Qed. 505 | Lemma rw_ren_rename_rename_comp l c b a r s : 506 | eqr (r_comp (r_comp l (r_rename c r b)) (r_rename b s a)) 507 | (r_comp l (r_rename c (r_comp r s) a)). 508 | Proof. easy_eqr. Qed. 509 | Lemma rw_ren_rename_shift c b r s : 510 | eqr (r_comp (r_rename c r b) (r_shift b s)) 511 | (r_shift c (r_comp r s)). 512 | Proof. easy_eqr. Qed. 513 | Lemma rw_ren_rename_shift_comp l c b r s : 514 | eqr (r_comp (r_comp l (r_rename c r b)) (r_shift b s)) 515 | (r_comp l (r_shift c (r_comp r s))). 516 | Proof. easy_eqr. Qed. 517 | Lemma rw_ren_subst_rename u r a s b : 518 | eqr (r_comp (r_subst u r a) (r_rename a s b)) 519 | (r_subst u (r_comp r s) b). 520 | Proof. easy_eqr. Qed. 521 | Lemma rw_ren_subst_rename_comp l u r a s b : 522 | eqr (r_comp (r_comp l (r_subst u r a)) (r_rename a s b)) 523 | (r_comp l (r_subst u (r_comp r s) b)). 524 | Proof. easy_eqr. Qed. 525 | Lemma rw_ren_subst_shift u r a s : 526 | eqr (r_comp (r_subst u r a) (r_shift a s)) 527 | (r_comp r s). 528 | Proof. 529 | repeat intro; simpl_term_eq. 530 | (* FIXME: this again. Should we remove free? 531 | Should we do this automatically? 532 | Should we define eqr in terms of all vars, 533 | not just free ones? *) 534 | repeat match goal with 535 | | |- context C [@free (S ?v) ?x] => 536 | change (@free (S v) x) with (@wkv v (free x)) end. 537 | autorewrite with rw_renaming; simpl_term_eq. 538 | autorewrite with rw_renaming; simpl_term_eq. 539 | Qed. 540 | Lemma rw_ren_subst_shift_comp l u r a s : 541 | eqr (r_comp (r_comp l (r_subst u r a)) (r_shift a s)) 542 | (r_comp l (r_comp r s)). 543 | Proof. 544 | repeat intro; simpl_term_eq. 545 | repeat match goal with 546 | | |- context C [@free (S ?v) ?x] => 547 | change (@free (S v) x) with (@wkv v (free x)) end. 548 | autorewrite with rw_renaming; simpl_term_eq. 549 | autorewrite with rw_renaming; simpl_term_eq. 550 | Qed. 551 | Lemma rw_ren_shift_any a r s : 552 | eqr (r_comp (r_shift a r) s) (r_shift a (r_comp r s)). 553 | Proof. 554 | repeat intro; simpl_term_eq. 555 | autorewrite with rw_renaming; simpl_term_eq. 556 | Qed. 557 | Lemma rw_ren_shift_any_comp l a r s : 558 | eqr (r_comp (r_comp l (r_shift a r)) s) 559 | (r_comp l (r_shift a (r_comp r s))). 560 | Proof. 561 | repeat intro; simpl_term_eq. 562 | autorewrite with rw_renaming; simpl_term_eq. 563 | Qed. 564 | Lemma rw_ren_any_subst r u s x : 565 | eqr (r_comp r (r_subst u s x)) 566 | (r_subst ([r] u) (r_comp r s) x). 567 | Proof. 568 | repeat intro; simpl_term_eq. 569 | destruct (compare_l0 t); simpl_term_eq; 570 | autorewrite with rw_renaming; simpl_term_eq. 571 | Qed. 572 | 573 | (* No _comp version required for any_subst *) 574 | (* Ensure crit pair of rw_ren_shift_any and 575 | rw_ren_any_subst is confluent *) 576 | Lemma rw_ren_shift_subst_crit a u r b : 577 | eqr (r_subst ([^a] u) (r_shift a r) b) 578 | (r_shift a (r_subst u r b)). 579 | Proof. 580 | Set Printing Implicit. 581 | repeat intro; simpl_term_eq. 582 | autorewrite with rw_renaming. 583 | repeat rewrite <- rw_apply_comp. 584 | repeat rewrite rw_apply_id. 585 | rewrite rw_push_bind_shift. 586 | repeat rewrite rw_apply_id. 587 | replace (bind u (@applyv r (S V) (closev b v))) 588 | with (@applyv (r_subst u r b) V v) 589 | by simpl_term_eq. 590 | replace (bind u (@applyv r (S (S V)) (closev b (wkv v)))) 591 | with (@applyv (r_subst u r b) (S V) (wkv v)) 592 | by simpl_term_eq. 593 | rewrite rw_applyv_wkv. 594 | simpl_term_eq. 595 | Qed. 596 | 597 | Hint Rewrite 598 | @rw_ren_id_comp 599 | @rw_ren_comp_id 600 | @rw_ren_comp_assoc 601 | @rw_ren_under_id 602 | @rw_ren_rename_rename 603 | @rw_ren_rename_rename_comp 604 | @rw_ren_rename_shift 605 | @rw_ren_rename_shift_comp 606 | @rw_ren_subst_rename 607 | @rw_ren_subst_rename_comp 608 | @rw_ren_subst_shift 609 | @rw_ren_subst_shift_comp 610 | @rw_ren_shift_any 611 | @rw_ren_shift_any_comp 612 | @rw_ren_any_subst 613 | @rw_ren_shift_subst_crit : rw_renaming. 614 | 615 | (* Commuting open and close with apply *) 616 | 617 | Lemma rw_push_open a r {V} (t : @term (S V)) : 618 | open a ([r] t) = [r,, a] (open a t). 619 | Proof. simpl_term_eq. Qed. 620 | Lemma rw_push_close a r {V} (t : @term V) : 621 | [r] (close a t) = close a ([r,, a] t). 622 | Proof. simpl_term_eq. Qed. 623 | Hint Rewrite @rw_push_open @rw_push_close : rw_renaming. 624 | 625 | Arguments term / {V}. 626 | 627 | Arguments unit {N} / V t. 628 | 629 | Arguments kleisli {N M} f V !t /. 630 | 631 | Ltac names := 632 | repeat progress 633 | (cbn; autorewrite with rw_renaming rw_vars rw_names). 634 | 635 | Tactic Notation "names" "in" hyp(H) := 636 | repeat progress 637 | (cbn in H; autorewrite with rw_renaming in H). 638 | 639 | Lemma swap_subst_subst_distinct {x y} 640 | (Hd : distinct_names x y) {u v R} : 641 | eqr (r_subst u (r_subst v R x) y) 642 | (r_subst v (r_subst u R y) x). 643 | Proof. 644 | intros V p. simpl_term_eq. 645 | rewrite (swap_close_close Hd). 646 | destruct (closev y (closev x _)) as [?|[|[|?]]]; 647 | names; simpl_term_eq. 648 | Qed. 649 | 650 | End Renamings. 651 | -------------------------------------------------------------------------------- /src/Var.v: -------------------------------------------------------------------------------- 1 | Require Import String Omega Morph. 2 | 3 | (* Name indices are [nat]s *) 4 | Definition index := nat. 5 | 6 | (* Shift up all indices greater than or equal to [i] *) 7 | Fixpoint shift_idx (i : index) (j : index) : index := 8 | match i with 9 | | 0 => S j 10 | | S i => 11 | match j with 12 | | 0 => 0 13 | | S j => S (shift_idx i j) 14 | end 15 | end. 16 | 17 | Lemma rw_shift_idx_ge i j : 18 | i <= j -> 19 | shift_idx i j = S j. 20 | Proof. 21 | generalize dependent j. 22 | induction i; destruct j; intros; try easy. 23 | cbn; auto with arith. 24 | Qed. 25 | 26 | Lemma rw_shift_idx_lt i j : 27 | S j <= i -> 28 | shift_idx i j = j. 29 | Proof. 30 | generalize dependent j. 31 | induction i; destruct j; intros; try easy. 32 | cbn; auto with arith. 33 | Qed. 34 | 35 | Lemma rw_shift_idx_same i : 36 | shift_idx i i = S i. 37 | Proof. 38 | apply rw_shift_idx_ge; auto. 39 | Qed. 40 | 41 | (* Partial inverse of [shift_idx]. Returns [None] iff [i = j]. *) 42 | Fixpoint inverse_shift_idx i j := 43 | match i with 44 | | 0 => 45 | match j with 46 | | 0 => None 47 | | S j => Some j 48 | end 49 | | S i => 50 | match j with 51 | | 0 => Some 0 52 | | S j => 53 | match inverse_shift_idx i j with 54 | | None => None 55 | | Some j' => Some (S j') 56 | end 57 | end 58 | end. 59 | 60 | Lemma rw_inverse_shift_idx_eq i j : 61 | i = j -> 62 | inverse_shift_idx i j = None. 63 | Proof. 64 | intros <-. 65 | induction i; try easy; cbn. 66 | rewrite IHi; easy. 67 | Qed. 68 | 69 | Lemma rw_inverse_shift_idx_lt i j : 70 | S j <= i -> 71 | inverse_shift_idx i j = Some j. 72 | Proof. 73 | generalize dependent j. 74 | induction i; destruct j; intros Heq; try easy; cbn. 75 | rewrite IHi; auto with arith. 76 | Qed. 77 | 78 | Lemma rw_inverse_shift_idx_gt i j : 79 | S i <= j -> 80 | inverse_shift_idx i j = Some (pred j). 81 | Proof. 82 | generalize dependent j. 83 | induction i; destruct j; intros Heq; try easy; cbn. 84 | rewrite IHi; f_equal; omega. 85 | Qed. 86 | 87 | Lemma rw_inverse_shift_idx_same i : 88 | inverse_shift_idx i i = None. 89 | Proof. 90 | apply rw_inverse_shift_idx_eq; easy. 91 | Qed. 92 | 93 | (* Tactics for handling the above operations *) 94 | 95 | Arguments shift_idx !i !j. 96 | Arguments inverse_shift_idx !i !j. 97 | 98 | (* Rewrite [shift_idx]s, [unshift_idx]s and [inverse_shift_idx]s where 99 | the order of the parameters can be determined by the omega tactic *) 100 | Ltac simpl_idxs := 101 | repeat progress 102 | (match goal with 103 | | |- context [shift_idx ?i ?j] => 104 | first 105 | [ rewrite (rw_shift_idx_ge i j) by omega 106 | | rewrite (rw_shift_idx_lt i j) by omega ] 107 | | |- context [inverse_shift_idx ?i ?j] => 108 | first 109 | [ rewrite (rw_inverse_shift_idx_eq i j) by omega 110 | | rewrite (rw_inverse_shift_idx_lt i j) by omega 111 | | rewrite (rw_inverse_shift_idx_gt i j) by omega ] 112 | end; cbn). 113 | 114 | (* Case split on the order of the parameters, then simplify any 115 | [shift_idx]s and [inverse_shift_idx]s affected by the 116 | ordering. *) 117 | Ltac case_order i j := 118 | destruct (Compare_dec.lt_eq_lt_dec i j) as [[|]|]; 119 | simpl_idxs. 120 | 121 | (* Inverses *) 122 | 123 | Lemma rw_inverse_shift_idx_shift_idx i j : 124 | inverse_shift_idx i (shift_idx i j) = Some j. 125 | Proof. 126 | case_order i j; easy. 127 | Qed. 128 | 129 | Lemma rw_shift_idx_inverse_shift_idx {i j} : 130 | option_map (shift_idx i) (inverse_shift_idx i j) 131 | = if Nat.eq_dec i j then None else Some j. 132 | Proof. 133 | case_order i j; 134 | destruct (Nat.eq_dec i j); f_equal; try omega. 135 | Qed. 136 | 137 | (* Comparison for free name indices *) 138 | 139 | Inductive index_comparison (i : index) : index -> Set := 140 | | same_idx : index_comparison i i 141 | | diff_idx j : index_comparison i (shift_idx i j). 142 | 143 | Fixpoint compare_idx (i : index) 144 | : forall j, index_comparison i j := 145 | match i with 146 | | 0 => fun j => 147 | match j with 148 | | 0 => same_idx _ 149 | | S j => diff_idx _ j 150 | end 151 | | S i => fun j => 152 | match j with 153 | | 0 => diff_idx _ 0 154 | | S j => 155 | match compare_idx i j with 156 | | same_idx _ => same_idx _ 157 | | diff_idx _ j' => diff_idx _ (S j') 158 | end 159 | end 160 | end. 161 | 162 | Lemma rw_compare_idx_same : 163 | forall i, compare_idx i i = same_idx i. 164 | Proof. induction i; cbn; try rewrite IHi; easy. Qed. 165 | 166 | Lemma rw_compare_idx_shift : 167 | forall i j, compare_idx i (shift_idx i j) = diff_idx i j. 168 | Proof. induction i; destruct j; cbn; try rewrite IHi; easy. Qed. 169 | 170 | Hint Rewrite @rw_shift_idx_same @rw_inverse_shift_idx_same 171 | @rw_inverse_shift_idx_shift_idx @rw_shift_idx_inverse_shift_idx 172 | @rw_compare_idx_same @rw_compare_idx_shift 173 | : rw_idxs. 174 | 175 | Arguments shift_idx : simpl never. 176 | Arguments inverse_shift_idx : simpl never. 177 | 178 | (* Free names are a pair of a string and an index *) 179 | 180 | Set Primitive Projections. 181 | Record name := mkname { n_string : string; n_index : index }. 182 | Add Printing Constructor name. 183 | Unset Primitive Projections. 184 | 185 | Definition name_of_string s := mkname s 0. 186 | Coercion name_of_string : string >-> name. 187 | Bind Scope string_scope with name. 188 | 189 | Lemma name_dec (a : name) (b : name) : 190 | {a = b} + {a <> b}. 191 | Proof. 192 | destruct (string_dec (n_string a) (n_string b)). 193 | - destruct (Nat.eq_dec (n_index a) (n_index b)). 194 | + left; destruct a, b; cbn in *; subst; easy. 195 | + right; intro; subst; contradiction. 196 | - right; intro; subst; contradiction. 197 | Qed. 198 | 199 | Definition indistinct_names a b := 200 | n_string a = n_string b. 201 | 202 | Definition distinct_names a b := 203 | not (indistinct_names a b). 204 | 205 | Lemma distinct_names_dec a b : 206 | { distinct_names a b } + { indistinct_names a b }. 207 | Proof. 208 | destruct (string_dec (n_string a) (n_string b)). 209 | - right; easy. 210 | - left; easy. 211 | Qed. 212 | 213 | Lemma distinct_names_same a : 214 | ~ (distinct_names a a). 215 | Proof. 216 | intro Hd. 217 | unfold distinct_names, indistinct_names in Hd. 218 | easy. 219 | Qed. 220 | 221 | Lemma distinct_names_symmetric {a b} : 222 | distinct_names a b -> 223 | distinct_names b a. 224 | Proof. 225 | unfold distinct_names; intuition. 226 | Qed. 227 | 228 | (* Shift the index of a name up by one *) 229 | Definition succ_name (a : name) := 230 | mkname (n_string a) (S (n_index a)). 231 | 232 | (* Shift up all names with the same string as [a] and an index 233 | greater than or equal to [a]'s *) 234 | Definition shift_name (a : name) (b : name) := 235 | mkname (n_string b) 236 | (if string_dec (n_string a) (n_string b) then 237 | shift_idx (n_index a) (n_index b) 238 | else 239 | n_index b). 240 | Arguments shift_name !a !b. 241 | 242 | Lemma rw_shift_name_distinct a b : 243 | distinct_names a b -> 244 | shift_name a b = b. 245 | Proof. 246 | unfold shift_name. 247 | destruct (string_dec (n_string a) (n_string b)); easy. 248 | Qed. 249 | 250 | Lemma rw_shift_name_indistinct a b : 251 | indistinct_names a b -> 252 | shift_name a b 253 | = mkname (n_string b) (shift_idx (n_index a) (n_index b)). 254 | Proof. 255 | unfold shift_name. 256 | destruct (string_dec (n_string a) (n_string b)); easy. 257 | Qed. 258 | 259 | Lemma rw_shift_name_same a : 260 | shift_name a a = succ_name a. 261 | Proof. 262 | rewrite rw_shift_name_indistinct by easy. 263 | rewrite rw_shift_idx_same; easy. 264 | Qed. 265 | 266 | (* Partial inverse of [shift_name]. Returns [None] iff [b = a]. *) 267 | Definition inverse_shift_name (a : name) (b : name) := 268 | if string_dec (n_string a) (n_string b) then 269 | option_map (mkname (n_string b)) 270 | (inverse_shift_idx (n_index a) (n_index b)) 271 | else 272 | Some b. 273 | 274 | Lemma rw_inverse_shift_name_distinct a b : 275 | distinct_names a b -> 276 | inverse_shift_name a b = Some b. 277 | Proof. 278 | unfold inverse_shift_name. 279 | destruct (string_dec (n_string a) (n_string b)); easy. 280 | Qed. 281 | 282 | Lemma rw_inverse_shift_name_indistinct a b : 283 | indistinct_names a b -> 284 | inverse_shift_name a b 285 | = option_map (mkname (n_string b)) 286 | (inverse_shift_idx (n_index a) (n_index b)). 287 | Proof. 288 | unfold inverse_shift_name. 289 | destruct (string_dec (n_string a) (n_string b)); easy. 290 | Qed. 291 | 292 | Lemma rw_inverse_shift_name_same a : 293 | inverse_shift_name a a = None. 294 | Proof. 295 | rewrite rw_inverse_shift_name_indistinct by easy. 296 | rewrite rw_inverse_shift_idx_same; easy. 297 | Qed. 298 | 299 | (* Rewrite [shift_idx]s, [unshift_idx]s and [inverse_shift_idx]s where 300 | the order of the parameters can be determined by the omega tactic *) 301 | Ltac simpl_names := 302 | repeat progress 303 | (match goal with 304 | | |- context [shift_name ?a ?b] => 305 | first 306 | [ rewrite (rw_shift_name_distinct a b) by easy 307 | | rewrite (rw_shift_name_indistinct a b) by easy ] 308 | | |- context [inverse_shift_name ?a ?b] => 309 | first 310 | [ rewrite (rw_inverse_shift_name_distinct a b) by easy 311 | | rewrite (rw_inverse_shift_name_indistinct a b) by easy ] 312 | end; cbn). 313 | 314 | (* Case split on the strings of the parameters, then simplify any 315 | [shift_names]s and [inverse_shift_names]s affected by the 316 | ordering. *) 317 | Ltac case_strings a b := 318 | destruct (distinct_names_dec a b); 319 | simpl_names. 320 | 321 | (* Inverses *) 322 | 323 | Lemma rw_inverse_shift_name_shift_name a b : 324 | inverse_shift_name a (shift_name a b) = Some b. 325 | Proof. 326 | case_strings a b; autorewrite with rw_idxs; easy. 327 | Qed. 328 | 329 | Lemma rw_shift_name_inverse_shift_name a b : 330 | option_map (shift_name a) (inverse_shift_name a b) 331 | = if distinct_names_dec a b then Some b 332 | else if Nat.eq_dec (n_index a) (n_index b) then None 333 | else Some b. 334 | Proof. 335 | case_strings a b; try easy; 336 | case_order (n_index a) (n_index b); 337 | simpl_names; simpl_idxs; 338 | try replace (S (pred (n_index b))) with (n_index b) by omega; 339 | destruct (Nat.eq_dec (n_index a) (n_index b)); try omega; easy. 340 | Qed. 341 | 342 | Hint Rewrite @rw_shift_name_same @rw_inverse_shift_name_same 343 | @rw_inverse_shift_name_shift_name 344 | : rw_names. 345 | 346 | (* Renaming operation on names *) 347 | 348 | Definition rename_name a b c := 349 | match inverse_shift_name b c with 350 | | Some c' => shift_name a c' 351 | | None => a 352 | end. 353 | 354 | Arguments rename_name !a !b !c. 355 | 356 | Lemma rw_rename_name_distinct a b c : 357 | distinct_names a c -> 358 | rename_name b a c = shift_name b c. 359 | Proof. 360 | intros. 361 | unfold rename_name. 362 | rewrite (rw_inverse_shift_name_distinct a c) by easy. 363 | easy. 364 | Qed. 365 | 366 | Lemma rw_rename_name_both_distinct a b c : 367 | distinct_names a c -> 368 | distinct_names b c -> 369 | rename_name b a c = c. 370 | Proof. 371 | intros. 372 | rewrite (rw_rename_name_distinct a b c) by easy. 373 | rewrite (rw_shift_name_distinct b c) by easy. 374 | easy. 375 | Qed. 376 | 377 | Lemma rw_rename_name_same a b : rename_name b a a = b. 378 | Proof. 379 | unfold rename_name. 380 | autorewrite with rw_names; easy. 381 | Qed. 382 | 383 | Lemma rw_rename_name_shift_name a b c : 384 | rename_name b a (shift_name a c) = shift_name b c. 385 | Proof. 386 | unfold rename_name. 387 | autorewrite with rw_names; easy. 388 | Qed. 389 | 390 | Lemma rw_rename_name_rename_name a b c d : 391 | rename_name b a (rename_name a c d) 392 | = rename_name b c d. 393 | Proof. 394 | unfold rename_name. 395 | case_strings c d; 396 | case_order (n_index c) (n_index d); 397 | autorewrite with rw_names; easy. 398 | Qed. 399 | 400 | Hint Rewrite @rw_rename_name_same @rw_rename_name_shift_name 401 | @rw_rename_name_rename_name 402 | : rw_names. 403 | 404 | (* Comparison for free names *) 405 | 406 | Inductive name_comparison (a : name) : name -> Set := 407 | | same_name : name_comparison a a 408 | | diff_name b : name_comparison a (shift_name a b). 409 | 410 | Lemma compare_names a b : name_comparison a b. 411 | Proof. 412 | destruct a as [an ai], b as [bn bi]; cbn. 413 | remember (string_dec an bn) as Hdec. 414 | destruct Hdec; subst. 415 | - destruct (compare_idx ai bi) as [|j]. 416 | + apply same_name. 417 | + replace (mkname bn (shift_idx ai j)) 418 | with (shift_name (mkname bn ai) (mkname bn j)); 419 | try apply diff_name. 420 | unfold shift_name, n_string. 421 | rewrite <- HeqHdec; easy. 422 | - replace (mkname bn bi) 423 | with (shift_name (mkname an ai) (mkname bn bi)); 424 | try apply diff_name. 425 | unfold shift_name, n_string. 426 | rewrite <- HeqHdec; easy. 427 | Qed. 428 | 429 | (* Bound variables are represented by a level *) 430 | 431 | Definition Zero := Empty_set. 432 | 433 | Inductive Succ {S : Set} : Set := l0 | lS (s : S). 434 | 435 | Fixpoint level (V : nat) : Set := 436 | match V with 437 | | 0 => Zero 438 | | S V => @Succ (level V) 439 | end. 440 | Arguments level !V. 441 | 442 | (* Variables are either free names or bound levels *) 443 | 444 | Inductive var {V : nat} := 445 | | free (name : name) 446 | | bound (l : level V). 447 | 448 | (* The core operations acting on variables *) 449 | 450 | Definition wkv {V} (v : @var V) : @var (S V) := 451 | match v with 452 | | free n => free n 453 | | bound l => @bound (S V) (lS l) 454 | end. 455 | 456 | Definition openv a {V} (v : @var (S V)) : @var V := 457 | match v with 458 | | free n => free (shift_name a n) 459 | | bound l0 => free a 460 | | bound (lS l) => @bound V l 461 | end. 462 | 463 | Definition closev a {V} (v : @var V) : @var (S V) := 464 | match v with 465 | | free n => 466 | match inverse_shift_name a n with 467 | | Some n' => free n' 468 | | None => @bound (S V) l0 469 | end 470 | | bound l => @bound (S V) (lS l) 471 | end. 472 | 473 | Definition bindv {V} (v : @var (S V)) : option (@var V) := 474 | match v with 475 | | free n => Some (free n) 476 | | bound l0 => None 477 | | bound (lS l) => Some (@bound V l) 478 | end. 479 | 480 | (* We don't want to reduce the operations if it just exposes 481 | the inner match. ([wkv] has no inner matches) *) 482 | 483 | Arguments openv : simpl nomatch. 484 | Arguments closev : simpl nomatch. 485 | Arguments bindv : simpl nomatch. 486 | 487 | (* Add reductions for [closev]. The other operations reduce 488 | directly by cbn. *) 489 | 490 | Lemma rw_closev_shift a b : 491 | forall {V}, @closev a V (free (shift_name a b)) = free b. 492 | Proof. 493 | unfold closev; autorewrite with rw_names; easy. 494 | Qed. 495 | 496 | Lemma rw_closev_same a : 497 | forall {V}, @closev a V (free a) = 498 | @bound (S V) (@l0 (level V)). 499 | Proof. 500 | unfold closev; autorewrite with rw_names; easy. 501 | Qed. 502 | 503 | Hint Rewrite @rw_closev_shift @rw_closev_same : rw_vars. 504 | 505 | (* Open and close on the same variable are inverses. Weaken 506 | is a right inverse of bind. *) 507 | 508 | Lemma rw_openv_closev a {V} (v : @var V) : 509 | openv a (closev a v) = v. 510 | Proof. 511 | destruct v as [n|l]; cbn; try easy. 512 | destruct (compare_names a n); autorewrite with rw_vars; easy. 513 | Qed. 514 | 515 | Lemma rw_closev_openv a {V} (v : @var (S V)) : 516 | closev a (openv a v) = v. 517 | Proof. 518 | destruct v as [n|[|l]]; cbn; 519 | autorewrite with rw_vars; easy. 520 | Qed. 521 | 522 | Lemma rw_bindv_wkv {V} (v : @var V) : 523 | bindv (wkv v) = Some v. 524 | Proof. destruct v; easy. Qed. 525 | 526 | Hint Rewrite @rw_closev_openv @rw_openv_closev @rw_bindv_wkv 527 | : rw_vars. 528 | 529 | (* [openv] and [closev] on distinct names *) 530 | Lemma closev_distinct {a b} (Hd : distinct_names a b) {V} : 531 | closev a (@free V b) = free b. 532 | Proof. 533 | unfold closev. 534 | rewrite (rw_inverse_shift_name_distinct a b) by easy. 535 | easy. 536 | Qed. 537 | 538 | Lemma openv_distinct {a b} (Hd : distinct_names a b) {V} : 539 | openv a (@free (S V) b) = free b. 540 | Proof. 541 | unfold openv. 542 | rewrite (rw_shift_name_distinct a b) by easy. 543 | easy. 544 | Qed. 545 | 546 | (* Combined operations *) 547 | 548 | Definition shiftv a {V} v := @openv a V (wkv v). 549 | Definition renv a b {V} v := @openv a V (closev b v). 550 | Definition substv a {V} v := @bindv V (closev a v). 551 | 552 | (* We want [shiftv] to reduce whenever [v] is a 553 | constructor, since [wkv] will always reduce in 554 | such cases. *) 555 | Arguments shiftv a {V} !v. 556 | 557 | (* Add reductions for [renv] and [substv] based on the 558 | similar reductions for [closev] *) 559 | 560 | Lemma rw_renv_free a b c : 561 | forall {V}, @renv b a V (free c) = 562 | free (rename_name b a c). 563 | Proof. 564 | intros; unfold renv. 565 | destruct (compare_names a c); 566 | autorewrite with rw_names rw_vars; easy. 567 | Qed. 568 | 569 | Lemma rw_substv_shift a c : 570 | forall {V}, @substv a V (free (shift_name a c)) = 571 | Some (free c). 572 | Proof. 573 | intros; unfold substv; autorewrite with rw_vars; easy. 574 | Qed. 575 | 576 | Lemma rw_substv_same a : 577 | forall {V}, @substv a V (free a) = None. 578 | Proof. 579 | intros; unfold substv; autorewrite with rw_vars; easy. 580 | Qed. 581 | 582 | Hint Rewrite @rw_renv_free @rw_substv_shift @rw_substv_same 583 | : rw_vars. 584 | 585 | (* Combined operation reductions based on identities *) 586 | 587 | Lemma rw_closev_shiftv a {V} (v : @var V) : 588 | closev a (shiftv a v) = wkv v. 589 | Proof. 590 | unfold shiftv; autorewrite with rw_vars; easy. 591 | Qed. 592 | 593 | Lemma rw_closev_renv a b {V} (v : @var V) : 594 | closev a (renv a b v) = closev b v. 595 | Proof. 596 | unfold renv; autorewrite with rw_vars; easy. 597 | Qed. 598 | 599 | Lemma rw_renv_same a {V} (v : @var V) : 600 | renv a a v = v. 601 | Proof. 602 | unfold renv; autorewrite with rw_vars; easy. 603 | Qed. 604 | 605 | Lemma rw_renv_openv a b {V} (v : @var (S V)) : 606 | renv b a (openv a v) = openv b v. 607 | Proof. 608 | unfold renv; autorewrite with rw_vars; easy. 609 | Qed. 610 | 611 | Lemma rw_renv_shiftv a b {V} (v : @var V) : 612 | renv b a (shiftv a v) = shiftv b v. 613 | Proof. 614 | unfold renv, shiftv; autorewrite with rw_vars; easy. 615 | Qed. 616 | 617 | Lemma rw_renv_renv a b c {V} (v : @var V) : 618 | renv b a (renv a c v) = renv b c v. 619 | Proof. 620 | unfold renv; autorewrite with rw_vars; easy. 621 | Qed. 622 | 623 | Lemma rw_substv_renv a b {V} (v : @var V) : 624 | substv a (renv a b v) = substv b v. 625 | Proof. 626 | unfold renv, substv; autorewrite with rw_vars; easy. 627 | Qed. 628 | 629 | Lemma rw_substv_shiftv a {V} (v : @var V) : 630 | substv a (shiftv a v) = Some v. 631 | Proof. 632 | unfold shiftv, substv; autorewrite with rw_vars; easy. 633 | Qed. 634 | 635 | Hint Rewrite @rw_closev_shiftv @rw_closev_renv @rw_renv_same 636 | @rw_renv_openv @rw_renv_shiftv @rw_renv_renv 637 | @rw_substv_renv @rw_substv_shiftv 638 | : rw_vars. 639 | 640 | (* Fold combined operations *) 641 | 642 | Lemma rw_shiftv_fold a {V} (v : @var V) : 643 | openv a (wkv v) = shiftv a v. 644 | Proof. easy. Qed. 645 | Lemma rw_renv_fold a b {V} (v : @var V) : 646 | openv b (closev a v) = renv b a v. 647 | Proof. easy. Qed. 648 | Lemma rw_substv_fold a {V} (v : @var V) : 649 | bindv (closev a v) = substv a v. 650 | Proof. easy. Qed. 651 | 652 | Hint Rewrite @rw_shiftv_fold @rw_renv_fold @rw_substv_fold 653 | : rw_vars. 654 | 655 | (* Combined operations on distinct names *) 656 | 657 | Lemma shiftv_distinct {a b} (Hd : distinct_names a b) {V} : 658 | shiftv a (@free V b) = free b. 659 | Proof. 660 | cbn; rewrite (rw_shift_name_distinct a b) by easy; easy. 661 | Qed. 662 | 663 | Lemma renv_distinct {a b c} (Hd : distinct_names a c) {V} : 664 | renv b a (@free V c) = free (shift_name b c). 665 | Proof. 666 | autorewrite with rw_vars. 667 | rewrite (rw_rename_name_distinct a b c) by easy; easy. 668 | Qed. 669 | 670 | Lemma renv_both_distinct {a b c} 671 | (Hd1 : distinct_names a c) (Hd2 : distinct_names b c) {V} : 672 | renv b a (@free V c) = free c. 673 | Proof. 674 | autorewrite with rw_vars. 675 | rewrite (rw_rename_name_both_distinct a b c) by easy; easy. 676 | Qed. 677 | 678 | Lemma substv_distinct {a b} (Hd : distinct_names a b) {V} : 679 | substv a (@free V b) = Some (free b). 680 | Proof. 681 | unfold substv; rewrite (closev_distinct Hd); easy. 682 | Qed. 683 | 684 | (* [wkv] commutes with [shiftv] and [renv]. We generally 685 | try to move [wkv] leftwards but carefully to avoid 686 | breaking confluence. *) 687 | 688 | Lemma swap_shiftv_wkv a {V} (v : @var V) : 689 | shiftv a (wkv v) = wkv (shiftv a v). 690 | Proof. destruct v; easy. Qed. 691 | 692 | Lemma swap_renv_wkv a b {V} (v : @var V) : 693 | renv a b (wkv v) = wkv (renv a b v). 694 | Proof. 695 | destruct v; cbn; autorewrite with rw_vars; easy. 696 | Qed. 697 | 698 | Lemma rw_bindv_shiftv_wkv a {V} (v : @var V) : 699 | bindv (shiftv a (wkv v)) = Some (shiftv a v). 700 | Proof. 701 | rewrite swap_shiftv_wkv; autorewrite with rw_vars; easy. 702 | Qed. 703 | 704 | Lemma rw_bindv_renv_wkv a b {V} (v : @var V) : 705 | bindv (renv a b (wkv v)) = Some (renv a b v). 706 | Proof. 707 | rewrite swap_renv_wkv; autorewrite with rw_vars; easy. 708 | Qed. 709 | 710 | Lemma rw_openv_shiftv_wkv a b {V} (v : @var V) : 711 | openv a (shiftv b (wkv v)) = shiftv a (shiftv b v). 712 | Proof. 713 | rewrite swap_shiftv_wkv; autorewrite with rw_vars; easy. 714 | Qed. 715 | 716 | Lemma rw_openv_renv_wkv a b c {V} (v : @var V) : 717 | openv a (renv b c (wkv v)) = shiftv a (renv b c v). 718 | Proof. 719 | rewrite swap_renv_wkv; autorewrite with rw_vars; easy. 720 | Qed. 721 | 722 | Lemma rw_closev_wkv_shiftv a {V} (v : @var V) : 723 | closev a (wkv (shiftv a v)) = wkv (wkv v). 724 | Proof. 725 | rewrite <- swap_shiftv_wkv; autorewrite with rw_vars; easy. 726 | Qed. 727 | 728 | Lemma rw_renv_wkv_shiftv a b {V} (v : @var V) : 729 | renv b a (wkv (shiftv a v)) = wkv (shiftv b v). 730 | Proof. 731 | rewrite swap_renv_wkv; autorewrite with rw_vars; easy. 732 | Qed. 733 | 734 | Lemma rw_substv_wkv_shiftv a {V} (v : @var V) : 735 | substv a (wkv (shiftv a v)) = Some (wkv v). 736 | Proof. 737 | rewrite <- swap_shiftv_wkv; autorewrite with rw_vars; easy. 738 | Qed. 739 | 740 | Lemma rw_closev_wkv_renv a b {V} (v : @var V) : 741 | closev a (wkv (renv a b v)) = closev b (wkv v). 742 | Proof. 743 | rewrite <- swap_renv_wkv; autorewrite with rw_vars; easy. 744 | Qed. 745 | 746 | Lemma rw_renv_wkv_renv a b c {V} (v : @var V) : 747 | renv b a (wkv (renv a c v)) = wkv (renv b c v). 748 | Proof. 749 | rewrite swap_renv_wkv; autorewrite with rw_vars; easy. 750 | Qed. 751 | 752 | Lemma rw_substv_wkv_renv a b {V} (v : @var V) : 753 | substv a (wkv (renv a b v)) = substv b (wkv v). 754 | Proof. 755 | rewrite <- swap_renv_wkv; autorewrite with rw_vars; easy. 756 | Qed. 757 | 758 | Lemma rw_shiftv_shiftv_wkv a b {V} (v : @var V) : 759 | shiftv a (shiftv b (wkv v)) = shiftv a (wkv (shiftv b v)). 760 | Proof. 761 | rewrite swap_shiftv_wkv; easy. 762 | Qed. 763 | 764 | Lemma rw_renv_shiftv_wkv a b c {V} (v : @var V) : 765 | renv b a (shiftv c (wkv v)) = renv b a (wkv (shiftv c v)). 766 | Proof. 767 | rewrite swap_shiftv_wkv; easy. 768 | Qed. 769 | 770 | Lemma rw_renv_renv_wkv a b c d {V} (v : @var V) : 771 | renv b a (renv c d (wkv v)) = renv b a (wkv (renv c d v)). 772 | Proof. 773 | rewrite swap_renv_wkv; easy. 774 | Qed. 775 | 776 | Hint Rewrite @rw_bindv_shiftv_wkv @rw_bindv_renv_wkv 777 | @rw_openv_shiftv_wkv @rw_openv_renv_wkv 778 | @rw_closev_wkv_shiftv @rw_renv_wkv_shiftv 779 | @rw_substv_wkv_shiftv @rw_closev_wkv_renv 780 | @rw_renv_wkv_renv @rw_substv_wkv_renv 781 | @rw_shiftv_shiftv_wkv @rw_renv_shiftv_wkv 782 | @rw_renv_renv_wkv 783 | : rw_vars. 784 | 785 | (* Comparison of vars *) 786 | 787 | Inductive var_comparison (a : name) {V} : @var V -> Set := 788 | | samev : var_comparison a (@free V a) 789 | | diffv v : var_comparison a (shiftv a v). 790 | 791 | Definition compare_vars a {V} (v : @var V) 792 | : var_comparison a v. 793 | destruct v as [b|l]. 794 | destruct (compare_names a b). 795 | - constructor. 796 | - change (free (shift_name a b)) 797 | with (shiftv a (@free V b)). 798 | constructor. 799 | - change (bound l) with (shiftv a (bound l)). 800 | constructor. 801 | Qed. 802 | 803 | (* Comparison with [bound l0] *) 804 | 805 | Inductive l0_comparison {V} : @var (S V) -> Set := 806 | | samel0 : l0_comparison (@bound (S V) (@l0 (level V))) 807 | | diffl0 v : l0_comparison (wkv v). 808 | 809 | Definition compare_l0 {V} (v : @var (S V)) : l0_comparison v. 810 | Proof. 811 | destruct v as [a|[|l]]. 812 | - change (@free (S V) a) 813 | with (@wkv V (free a)); constructor. 814 | - constructor. 815 | - change (@bound (S V) (lS l)) 816 | with (@wkv V (bound l)); constructor. 817 | Qed. 818 | 819 | (* A couple of useful commuting lemmas *) 820 | 821 | Lemma swap_shiftv_shiftv_distinct {a b} (Hd : distinct_names a b) 822 | {V} (v : @var V) : 823 | shiftv a (shiftv b v) = shiftv b (shiftv a v). 824 | Proof. 825 | destruct v as [c|l]; cbn; try easy. 826 | case_strings b c; try easy. 827 | rewrite (rw_shift_name_distinct a c). 828 | - rewrite (rw_shift_name_distinct a _); 829 | replace (n_string c) with (n_string b); easy. 830 | - unfold distinct_names, indistinct_names. 831 | replace (n_string c) with (n_string b); easy. 832 | Qed. 833 | 834 | Lemma swap_shift_close {a b} (Hd : distinct_names b a) 835 | {V} (v : @var V) : 836 | closev a (shiftv b v) = shiftv b (closev a v). 837 | Proof. 838 | destruct (compare_vars a v). 839 | - rewrite shiftv_distinct by easy. 840 | autorewrite with rw_vars rw_names. 841 | easy. 842 | - rewrite swap_shiftv_shiftv_distinct by easy. 843 | autorewrite with rw_vars rw_names. 844 | rewrite swap_shiftv_wkv; easy. 845 | Qed. 846 | 847 | Definition swap_bound {V} (v : @var (S (S V))) : @var (S (S V)) := 848 | match v with 849 | | free a => free a 850 | | bound l => 851 | @bound (S (S V)) 852 | (match l with 853 | | l0 => lS l0 854 | | lS l0 => l0 855 | | lS (lS v) => lS (lS v) 856 | end) 857 | end. 858 | 859 | Lemma swap_close_close {x y} 860 | (Hd : distinct_names x y) {V} {v : @var V} : 861 | closev x (closev y v) = swap_bound (closev y (closev x v)). 862 | Proof. 863 | destruct (compare_vars y v); autorewrite with rw_vars; cbn. 864 | - replace (free y) with (shiftv x (@free V y)). 865 | + autorewrite with rw_vars; easy. 866 | + rewrite (shiftv_distinct Hd); easy. 867 | - rewrite swap_shift_close 868 | by auto using distinct_names_symmetric. 869 | autorewrite with rw_vars. 870 | destruct v as [n|l]; cbn; try easy. 871 | destruct (compare_names x n); 872 | autorewrite with rw_vars; easy. 873 | Qed. 874 | 875 | (* Algebra of operations on [var 0] *) 876 | Inductive renaming {trm : Set} := 877 | | r_id 878 | | r_comp (r : renaming) (s : renaming) 879 | | r_shift (b : name) (r : renaming) 880 | | r_rename (b : name) (r : renaming) (a : name) 881 | | r_subst (t : trm) (r : renaming) (a : name). 882 | 883 | Declare Scope ren_scope. 884 | Notation "r1 ; r2" := (r_comp r1 r2) 885 | (at level 57, right associativity) : ren_scope. 886 | Notation "r ,, ^ a" := (r_shift a r) 887 | (at level 47, left associativity) : ren_scope. 888 | Notation "r ,, a <- b" := (r_rename a r b) 889 | (at level 47, left associativity, a at next level) : ren_scope. 890 | Notation "r ,, u // a" := (r_subst u r a) 891 | (at level 47, left associativity, u at next level) : ren_scope. 892 | Notation "^ a" := (r_shift a r_id) 893 | (at level 47, left associativity) : ren_scope. 894 | Notation "a <- b" := (r_rename a r_id b) 895 | (at level 47, left associativity) : ren_scope. 896 | Notation "u // a" := (r_subst u r_id a) 897 | (at level 47, left associativity) : ren_scope. 898 | Notation "r ,, a" := (r_rename a r a) 899 | (at level 47, left associativity) : ren_scope. 900 | 901 | Delimit Scope ren_scope with ren. 902 | 903 | Fixpoint total {trm : Set} (r : @renaming trm) : Prop := 904 | match r with 905 | | r_id => True 906 | | r_comp r s => total r /\ total s 907 | | r_shift b r => total r 908 | | r_rename b r a => total r 909 | | r_subst u r a => False 910 | end. 911 | 912 | Definition proj1 {A B : Prop} (H : A /\ B) := let (a, _) := H in a. 913 | Definition proj2 {A B : Prop} (H : A /\ B) := let (_, b) := H in b. 914 | 915 | Fixpoint applyt {trm : Set} (r : @renaming trm) : 916 | forall (rn : total r), morph (@var) 0 (@var) 0 := 917 | match r with 918 | | r_id => 919 | fun _ _ v => v 920 | | r_comp r s => 921 | fun rn V v => applyt r (proj1 rn) V (applyt s (proj2 rn) V v) 922 | | r_shift b r => 923 | fun rn V v => openv b (applyt r rn (S V) (wkv v)) 924 | | r_rename b r a => 925 | fun rn V v => openv b (applyt r rn (S V) (closev a v)) 926 | | r_subst _ _ _ => False_rec _ 927 | end. 928 | 929 | Arguments applyt {trm} !r rn {V} t /. 930 | 931 | Lemma rw_applyt_bound trm (r : @renaming trm) rn : 932 | forall {V} (v : level V), 933 | applyt r rn (bound v) = bound v. 934 | Proof. 935 | induction r; try destruct rn; cbn; intros; auto; 936 | repeat match goal with [ H : _ |- _ ] => rewrite H; clear H end; 937 | cbn; easy. 938 | Qed. 939 | 940 | Hint Rewrite @rw_applyt_bound : rw_names. 941 | 942 | Lemma rw_applyt_wkv : 943 | forall trm (r : @renaming trm) (rn : total r) {V} (v : @var V), 944 | applyt r rn (wkv v) = wkv (applyt r rn v). 945 | Proof. 946 | induction r; cbn; intuition. 947 | - rewrite IHr2, IHr1; reflexivity. 948 | - repeat rewrite IHr. 949 | autorewrite with rw_vars. 950 | rewrite <- swap_shiftv_wkv; reflexivity. 951 | - destruct (compare_vars a v). 952 | + autorewrite with rw_vars rw_names; reflexivity. 953 | + autorewrite with rw_vars. 954 | repeat rewrite IHr. 955 | autorewrite with rw_vars. 956 | rewrite <- swap_shiftv_wkv; reflexivity. 957 | Qed. 958 | 959 | Hint Rewrite @rw_applyt_wkv : rw_names. 960 | 961 | Lemma rw_applyt_wkv_free trm (r : @renaming trm) rn {a V} : 962 | applyt r rn (@free (S V) a) = wkv (applyt r rn (@free V a)). 963 | Proof. 964 | change (@free (S V) a) with (wkv (@free V a)). 965 | apply rw_applyt_wkv. 966 | Qed. 967 | 968 | Hint Rewrite @rw_applyt_wkv_free : rw_names. 969 | 970 | -------------------------------------------------------------------------------- /src/_CoqProject: -------------------------------------------------------------------------------- 1 | -R . SNames --------------------------------------------------------------------------------