diff options
| author | Rutger Broekhoff | 2024-06-26 20:50:18 +0200 |
|---|---|---|
| committer | Rutger Broekhoff | 2024-06-26 20:50:18 +0200 |
| commit | 73df1945b31c0beee88cf4476df4ccd09d31403b (patch) | |
| tree | ed00db26b711442e643f38b66888a3df56e33ebd /semprop.v | |
| download | mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.tar.gz mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.zip | |
Import Coq project
Diffstat (limited to 'semprop.v')
| -rw-r--r-- | semprop.v | 426 |
1 files changed, 426 insertions, 0 deletions
diff --git a/semprop.v b/semprop.v new file mode 100644 index 0000000..3af718d --- /dev/null +++ b/semprop.v | |||
| @@ -0,0 +1,426 @@ | |||
| 1 | Require Import Coq.Strings.String. | ||
| 2 | From stdpp Require Import gmap option tactics. | ||
| 3 | From mininix Require Import binop expr maptools matching relations sem. | ||
| 4 | |||
| 5 | Lemma ctx_item_inj E uf_ext uf_int : | ||
| 6 | is_ctx_item uf_ext uf_int E → | ||
| 7 | Inj (=) (=) E. | ||
| 8 | Proof. | ||
| 9 | intros Hctx. | ||
| 10 | destruct Hctx; intros e1' e2' Hinj; try congruence. | ||
| 11 | injection Hinj as Hinj. | ||
| 12 | by apply map_insert_rev_1_L in Hinj. | ||
| 13 | Qed. | ||
| 14 | |||
| 15 | Global Instance id_inj {A} : Inj (=) (=) (@id A). | ||
| 16 | Proof. by unfold Inj. Qed. | ||
| 17 | |||
| 18 | Lemma ctx_inj E uf_ext uf_int : | ||
| 19 | is_ctx uf_ext uf_int E → | ||
| 20 | Inj (=) (=) E. | ||
| 21 | Proof. | ||
| 22 | intros Hctx. | ||
| 23 | induction Hctx. | ||
| 24 | - apply id_inj. | ||
| 25 | - apply ctx_item_inj in H. | ||
| 26 | by apply compose_inj with (=). | ||
| 27 | Qed. | ||
| 28 | |||
| 29 | Lemma base_step_deterministic : deterministic base_step. | ||
| 30 | Proof. | ||
| 31 | unfold deterministic. | ||
| 32 | intros e0 e1 e2 H1 H2. | ||
| 33 | inv H1; inv H2; try done. | ||
| 34 | - destruct ys, ys0. by simplify_eq/=. | ||
| 35 | - destruct ys. by simplify_eq/=. | ||
| 36 | - destruct ys. by simplify_eq/=. | ||
| 37 | - apply map_insert_inv in H0. by destruct H0. | ||
| 38 | - destruct ys. by simplify_eq/=. | ||
| 39 | - destruct ys. by simplify_eq/=. | ||
| 40 | - destruct ys, ys0. by simplify_eq/=. | ||
| 41 | - exfalso. apply H3. by exists bs. | ||
| 42 | - exfalso. apply H. by exists bs. | ||
| 43 | - f_equal. by eapply matches_deterministic. | ||
| 44 | - apply map_insert_inv in H0 as [H01 H02]. | ||
| 45 | by simplify_eq /=. | ||
| 46 | - f_equal. by eapply binop_deterministic. | ||
| 47 | - rewrite H4 in H. by apply is_Some_None in H. | ||
| 48 | - exfalso. apply H4. by exists bs. | ||
| 49 | - rewrite H in H4. by apply is_Some_None in H4. | ||
| 50 | - exfalso. apply H. by exists bs. | ||
| 51 | Qed. | ||
| 52 | |||
| 53 | Reserved Notation "e '-/' uf '/->' e'" (right associativity, at level 55). | ||
| 54 | |||
| 55 | Inductive fstep : bool → expr → expr → Prop := | ||
| 56 | | F_Ctx e1 e2 E uf_ext uf_int : | ||
| 57 | is_ctx uf_ext uf_int E → | ||
| 58 | e1 -->base e2 → | ||
| 59 | E e1 -/uf_ext/-> E e2 | ||
| 60 | where "e -/ uf /-> e'" := (fstep uf e e'). | ||
| 61 | |||
| 62 | Hint Constructors fstep : core. | ||
| 63 | |||
| 64 | Lemma ufstep e e' : e -/false/-> e' ↔ e --> e'. | ||
| 65 | Proof. split; intros; inv H; by econstructor. Qed. | ||
| 66 | |||
| 67 | Lemma fstep_strongly_confluent_aux x uf y1 y2 : | ||
| 68 | x -/uf/-> y1 → | ||
| 69 | x -/uf/-> y2 → | ||
| 70 | y1 = y2 ∨ ∃ z, y1 -/uf/-> z ∧ y2 -/uf/-> z. | ||
| 71 | Proof. | ||
| 72 | intros Hstep1 Hstep2. | ||
| 73 | inv Hstep1 as [b11 b12 E1 uf_int uf_int1 Hctx1 Hbase1]. | ||
| 74 | inv Hstep2 as [b21 b22 E2 uf_int uf_int2 Hctx2 Hbase2]. | ||
| 75 | revert b11 b12 b21 b22 E2 Hbase1 Hbase2 Hctx2 H0. | ||
| 76 | induction Hctx1 as [|E1 E0]; | ||
| 77 | intros b11 b12 b21 b22 E2 Hbase1 Hbase2 Hctx2 H2; inv Hctx2. | ||
| 78 | - (* L: id, R: id *) left. by eapply base_step_deterministic. | ||
| 79 | - (* L: id, R: ∘ *) simplify_eq/=. | ||
| 80 | inv H. | ||
| 81 | + inv Hbase1. | ||
| 82 | * inv H0. | ||
| 83 | -- inv Hbase2. | ||
| 84 | -- inv H. | ||
| 85 | * inv H0. | ||
| 86 | -- inv Hbase2. | ||
| 87 | -- inv H1. inv H. | ||
| 88 | + inv Hbase1. | ||
| 89 | * inv H0. | ||
| 90 | -- inv Hbase2. | ||
| 91 | -- inv H. | ||
| 92 | * inv H0. | ||
| 93 | -- inv Hbase2. | ||
| 94 | -- inv H. | ||
| 95 | + inv Hbase1. | ||
| 96 | * inv H0. | ||
| 97 | -- inv Hbase2. | ||
| 98 | -- inv H. | ||
| 99 | * inv H0. | ||
| 100 | -- inv Hbase2. | ||
| 101 | -- inv H1. | ||
| 102 | + inv Hbase1. | ||
| 103 | * inv H0. | ||
| 104 | -- inv Hbase2. | ||
| 105 | -- inv H. | ||
| 106 | * inv H0. | ||
| 107 | -- inv Hbase2. | ||
| 108 | -- inv H1. | ||
| 109 | * inv H0. | ||
| 110 | -- inv Hbase2. | ||
| 111 | -- inv H1. inv H. | ||
| 112 | + inv Hbase1. | ||
| 113 | inv H0. | ||
| 114 | * inv Hbase2. | ||
| 115 | * inv H. | ||
| 116 | + inv Hbase1. | ||
| 117 | * inv H0. | ||
| 118 | -- inv Hbase2. | ||
| 119 | -- inv H. | ||
| 120 | * inv H0. | ||
| 121 | -- inv Hbase2. | ||
| 122 | -- inv H. | ||
| 123 | + inv Hbase1. | ||
| 124 | inv H0. | ||
| 125 | * inv Hbase2. | ||
| 126 | * inv H. | ||
| 127 | + inv Hbase1. | ||
| 128 | inv H0. | ||
| 129 | * inv Hbase2. | ||
| 130 | * inv H. | ||
| 131 | + inv Hbase1. | ||
| 132 | inv H0. | ||
| 133 | * inv Hbase2. | ||
| 134 | * inv H. | ||
| 135 | + inv Hbase1. | ||
| 136 | * inv H0. | ||
| 137 | -- inv Hbase2. | ||
| 138 | -- inv H1. | ||
| 139 | * inv H0. | ||
| 140 | -- inv Hbase2. | ||
| 141 | -- inv H1. | ||
| 142 | * inv H0. | ||
| 143 | -- inv Hbase2. | ||
| 144 | -- inv H1. | ||
| 145 | + inv Hbase1. | ||
| 146 | inv H0. | ||
| 147 | * inv Hbase2. | ||
| 148 | * inv H. | ||
| 149 | simplify_option_eq. | ||
| 150 | destruct sv; try discriminate. | ||
| 151 | exfalso. clear uf. simplify_option_eq. | ||
| 152 | apply fmap_insert_lookup in H1. simplify_option_eq. | ||
| 153 | revert bs0 x H H1 Heqo. | ||
| 154 | induction H2; intros bs0 x H' H1 Heqo. | ||
| 155 | -- inv Hbase2. | ||
| 156 | -- simplify_option_eq. | ||
| 157 | inv H. destruct H'; try discriminate. | ||
| 158 | inv H1. apply fmap_insert_lookup in H0. simplify_option_eq. | ||
| 159 | by eapply IHis_ctx. | ||
| 160 | + inv Hbase1. | ||
| 161 | - (* L: ∘, R: id *) | ||
| 162 | simplify_eq/=. | ||
| 163 | inv H. | ||
| 164 | + inv Hbase2. | ||
| 165 | * inv Hctx1. | ||
| 166 | -- inv Hbase1. | ||
| 167 | -- inv H. | ||
| 168 | * inv Hctx1. | ||
| 169 | -- inv Hbase1. | ||
| 170 | -- inv H0. inv H. | ||
| 171 | + inv Hbase2. | ||
| 172 | * inv Hctx1. | ||
| 173 | -- inv Hbase1. | ||
| 174 | -- inv H. | ||
| 175 | * inv Hctx1. | ||
| 176 | -- inv Hbase1. | ||
| 177 | -- inv H. | ||
| 178 | + inv Hbase2. | ||
| 179 | * inv Hctx1. | ||
| 180 | -- inv Hbase1. | ||
| 181 | -- inv H. | ||
| 182 | * inv Hctx1. | ||
| 183 | -- inv Hbase1. | ||
| 184 | -- inv H0. | ||
| 185 | + inv Hbase2. | ||
| 186 | * inv Hctx1. | ||
| 187 | -- inv Hbase1. | ||
| 188 | -- inv H. | ||
| 189 | * inv Hctx1. | ||
| 190 | -- inv Hbase1. | ||
| 191 | -- inv H0. | ||
| 192 | * inv Hctx1. | ||
| 193 | -- inv Hbase1. | ||
| 194 | -- inv H0. inv H. | ||
| 195 | + inv Hbase2. | ||
| 196 | inv Hctx1. | ||
| 197 | * inv Hbase1. | ||
| 198 | * inv H. | ||
| 199 | + inv Hbase2. | ||
| 200 | * inv Hctx1. | ||
| 201 | -- inv Hbase1. | ||
| 202 | -- inv H. | ||
| 203 | * inv Hctx1. | ||
| 204 | -- inv Hbase1. | ||
| 205 | -- inv H. | ||
| 206 | + inv Hbase2. | ||
| 207 | inv Hctx1. | ||
| 208 | * inv Hbase1. | ||
| 209 | * inv H. | ||
| 210 | + inv Hbase2. | ||
| 211 | inv Hctx1. | ||
| 212 | * inv Hbase1. | ||
| 213 | * inv H. | ||
| 214 | + inv Hbase2. | ||
| 215 | inv Hctx1. | ||
| 216 | * inv Hbase1. | ||
| 217 | * inv H. | ||
| 218 | + inv Hbase2. | ||
| 219 | * inv Hctx1. | ||
| 220 | -- inv Hbase1. | ||
| 221 | -- inv H0. | ||
| 222 | * inv Hctx1. | ||
| 223 | -- inv Hbase1. | ||
| 224 | -- inv H0. | ||
| 225 | * inv Hctx1. | ||
| 226 | -- inv Hbase1. | ||
| 227 | -- inv H0. | ||
| 228 | + inv Hbase2. | ||
| 229 | inv Hctx1. | ||
| 230 | * inv Hbase1. | ||
| 231 | * clear IHHctx1. | ||
| 232 | inv H. | ||
| 233 | simplify_option_eq. | ||
| 234 | destruct sv; try discriminate. | ||
| 235 | exfalso. simplify_option_eq. | ||
| 236 | apply fmap_insert_lookup in H0. simplify_option_eq. | ||
| 237 | revert bs0 x H H0 Heqo. | ||
| 238 | induction H1; intros bs0 x H' H0 Heqo. | ||
| 239 | -- inv Hbase1. | ||
| 240 | -- simplify_option_eq. | ||
| 241 | inv H. destruct H'; try discriminate. | ||
| 242 | inv H0. apply fmap_insert_lookup in H2. simplify_option_eq. | ||
| 243 | eapply IHis_ctx. | ||
| 244 | ++ apply H2. | ||
| 245 | ++ apply Heqo0. | ||
| 246 | + inv Hbase2. | ||
| 247 | - (* L: ∘, R: ∘ *) | ||
| 248 | simplify_option_eq. | ||
| 249 | inv H; inv H0. | ||
| 250 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 251 | as [IH | [z [IHz1 IHz2]]]. | ||
| 252 | * left. congruence. | ||
| 253 | * right. exists (X_Select z xs). | ||
| 254 | split; [inv IHz1 | inv IHz2]; | ||
| 255 | eapply F_Ctx with (E := (λ e, X_Select e xs) ∘ E); try done; | ||
| 256 | by eapply IsCtx_Compose. | ||
| 257 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 258 | as [IH | [z [IHz1 IHz2]]]. | ||
| 259 | * left. congruence. | ||
| 260 | * right. exists (X_SelectOr z xs e2). | ||
| 261 | split; [inv IHz1 | inv IHz2]; | ||
| 262 | eapply F_Ctx with (E := (λ e1, X_SelectOr e1 xs e2) ∘ E); try done; | ||
| 263 | by eapply IsCtx_Compose. | ||
| 264 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 265 | as [IH | [z [IHz1 IHz2]]]. | ||
| 266 | * left. congruence. | ||
| 267 | * right. exists (X_Incl z e2). | ||
| 268 | split; [inv IHz1 | inv IHz2]; | ||
| 269 | eapply F_Ctx with (E := (λ e1, X_Incl e1 e2) ∘ E); try done; | ||
| 270 | by eapply IsCtx_Compose. | ||
| 271 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 272 | as [IH | [z [IHz1 IHz2]]]. | ||
| 273 | * left. congruence. | ||
| 274 | * right. exists (X_Apply z e2). | ||
| 275 | split; [inv IHz1 | inv IHz2]; | ||
| 276 | eapply F_Ctx with (E := (λ e1, X_Apply e1 e2) ∘ E); try done; | ||
| 277 | by eapply IsCtx_Compose. | ||
| 278 | + inv Hctx1; simplify_option_eq. | ||
| 279 | * inv Hbase1. | ||
| 280 | * inv H. | ||
| 281 | + inv H1; simplify_option_eq. | ||
| 282 | * inv Hbase2. | ||
| 283 | * inv H. | ||
| 284 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H0) | ||
| 285 | as [IH | [z [IHz1 IHz2]]]. | ||
| 286 | * left. congruence. | ||
| 287 | * right. exists (X_Apply (V_AttrsetFn m e1) z). | ||
| 288 | split; [inv IHz1 | inv IHz2]; | ||
| 289 | eapply F_Ctx with (E := (λ e2, X_Apply (V_AttrsetFn m e1) e2) ∘ E); | ||
| 290 | try done; | ||
| 291 | by eapply IsCtx_Compose. | ||
| 292 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 293 | as [IH | [z [IHz1 IHz2]]]. | ||
| 294 | * left. congruence. | ||
| 295 | * right. exists (X_Cond z e2 e3). | ||
| 296 | split; [inv IHz1 | inv IHz2]; | ||
| 297 | eapply F_Ctx with (E := (λ e1, X_Cond e1 e2 e3) ∘ E); try done; | ||
| 298 | by eapply IsCtx_Compose. | ||
| 299 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 300 | as [IH | [z [IHz1 IHz2]]]. | ||
| 301 | * left. congruence. | ||
| 302 | * right. exists (X_Assert z e2). | ||
| 303 | split; [inv IHz1 | inv IHz2]; | ||
| 304 | eapply F_Ctx with (E := (λ e1, X_Assert e1 e2) ∘ E); try done; | ||
| 305 | by eapply IsCtx_Compose. | ||
| 306 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H) | ||
| 307 | as [IH | [z [IHz1 IHz2]]]. | ||
| 308 | * left. congruence. | ||
| 309 | * right. exists (X_Op op z e2). | ||
| 310 | split; [inv IHz1 | inv IHz2]; | ||
| 311 | eapply F_Ctx with (E := (λ e1, X_Op op e1 e2) ∘ E); try done; | ||
| 312 | by eapply IsCtx_Compose. | ||
| 313 | + inv Hctx1; simplify_option_eq. | ||
| 314 | * inv Hbase1. | ||
| 315 | * inv H0. | ||
| 316 | + inv H1; simplify_option_eq. | ||
| 317 | * inv Hbase2. | ||
| 318 | * inv H0. | ||
| 319 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H0) | ||
| 320 | as [IH | [z [IHz1 IHz2]]]. | ||
| 321 | * left. congruence. | ||
| 322 | * right. exists (X_Op op v1 z). | ||
| 323 | split; [inv IHz1 | inv IHz2]; | ||
| 324 | eapply F_Ctx with (E := (λ e2, X_Op op v1 e2) ∘ E); try done; | ||
| 325 | by eapply IsCtx_Compose. | ||
| 326 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 327 | as [IH | [z [IHz1 IHz2]]]. | ||
| 328 | * left. congruence. | ||
| 329 | * right. exists (X_HasAttr z x). | ||
| 330 | split; [inv IHz1 | inv IHz2]; | ||
| 331 | eapply F_Ctx with (E := (λ e, X_HasAttr e x) ∘ E); try done; | ||
| 332 | by eapply IsCtx_Compose. | ||
| 333 | + destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H2) | ||
| 334 | as [IH | [z [IHz1 IHz2]]]. | ||
| 335 | * left. congruence. | ||
| 336 | * right. exists (X_Force z). | ||
| 337 | split. | ||
| 338 | -- inv IHz1. | ||
| 339 | (* apply is_ctx_uf_false_impl_true in H0 as []. *) | ||
| 340 | eapply F_Ctx with (E := X_Force ∘ E); try done; | ||
| 341 | by eapply IsCtx_Compose. | ||
| 342 | -- inv IHz2. | ||
| 343 | eapply F_Ctx with (E := X_Force ∘ E); try done; | ||
| 344 | by eapply IsCtx_Compose. | ||
| 345 | + destruct (decide (x0 = x)). | ||
| 346 | * simplify_option_eq. | ||
| 347 | apply map_insert_rev_L in H2 as [H21 H22]. | ||
| 348 | destruct (IHHctx1 _ _ _ _ E4 Hbase1 Hbase2 H1 H21) | ||
| 349 | as [IH | [z [IHz1 IHz2]]]. | ||
| 350 | -- left. rewrite IH. do 2 f_equal. | ||
| 351 | by apply map_insert_L. | ||
| 352 | -- right. exists (V_Attrset (<[x := z]>bs)). | ||
| 353 | split. | ||
| 354 | ++ inv IHz1. | ||
| 355 | eapply F_Ctx | ||
| 356 | with (E := (λ e, X_V $ V_Attrset (<[x := e]>bs)) ∘ E); try done. | ||
| 357 | by eapply IsCtx_Compose. | ||
| 358 | ++ inv IHz2. | ||
| 359 | rewrite (map_insert_L _ _ _ _ _ (eq_refl (E e1)) H22). | ||
| 360 | eapply F_Ctx | ||
| 361 | with (E := (λ e, X_V $ V_Attrset (<[x := e]>bs)) ∘ E); try done. | ||
| 362 | by eapply IsCtx_Compose. | ||
| 363 | * simplify_option_eq. | ||
| 364 | right. exists (V_Attrset (<[x := E0 b12]>(<[x0 := E4 b22]>bs))). | ||
| 365 | split. | ||
| 366 | -- apply map_insert_ne_inv in H2 as H3; try done. | ||
| 367 | destruct H3. | ||
| 368 | replace bs with (<[x0 := E4 b21]>bs) at 1; try by rewrite insert_id. | ||
| 369 | setoid_rewrite insert_commute; try by congruence. | ||
| 370 | eapply F_Ctx | ||
| 371 | with (E := (λ e, X_V $ V_Attrset | ||
| 372 | (<[x0 := e]>(<[x := E0 b12]> bs))) ∘ E4). | ||
| 373 | ++ by eapply IsCtx_Compose. | ||
| 374 | ++ done. | ||
| 375 | -- apply map_insert_ne_inv in H2 as H3; try done. | ||
| 376 | destruct H3 as [H31 [H32 H33]]. | ||
| 377 | replace bs0 with (<[x := E0 b11]>bs0) at 1; | ||
| 378 | try by rewrite insert_id. | ||
| 379 | rewrite insert_commute; try by congruence. | ||
| 380 | replace (<[x:=E0 b11]> (<[x0:=E4 b22]> bs0)) | ||
| 381 | with (<[x:=E0 b11]> (<[x0:=E4 b22]> bs)). | ||
| 382 | 2: { | ||
| 383 | rewrite <-insert_delete_insert. | ||
| 384 | setoid_rewrite <-insert_delete_insert at 2. | ||
| 385 | rewrite delete_insert_ne by congruence. | ||
| 386 | rewrite delete_commute. rewrite <-H33. | ||
| 387 | rewrite insert_delete_insert. | ||
| 388 | rewrite <-delete_insert_ne by congruence. | ||
| 389 | by rewrite insert_delete_insert. | ||
| 390 | } | ||
| 391 | eapply F_Ctx with (E := (λ e, X_V $ V_Attrset | ||
| 392 | (<[x := e]>(<[x0 := E4 b22]>bs))) ∘ E0). | ||
| 393 | ++ by eapply IsCtx_Compose. | ||
| 394 | ++ done. | ||
| 395 | Qed. | ||
| 396 | |||
| 397 | Lemma step_strongly_confluent_aux x y1 y2 : | ||
| 398 | x --> y1 → | ||
| 399 | x --> y2 → | ||
| 400 | y1 = y2 ∨ ∃ z, y1 --> z ∧ y2 --> z. | ||
| 401 | Proof. | ||
| 402 | intros Hstep1 Hstep2. | ||
| 403 | apply ufstep in Hstep1, Hstep2. | ||
| 404 | destruct (fstep_strongly_confluent_aux _ _ _ _ Hstep1 Hstep2) as [|[?[]]]. | ||
| 405 | - by left. | ||
| 406 | - right. exists x0. split; by apply ufstep. | ||
| 407 | Qed. | ||
| 408 | |||
| 409 | Theorem step_strongly_confluent : strongly_confluent step. | ||
| 410 | Proof. | ||
| 411 | intros x y1 y2 Hy1 Hy2. | ||
| 412 | destruct (step_strongly_confluent_aux x y1 y2 Hy1 Hy2) as [|[z [Hz1 Hz2]]]. | ||
| 413 | - exists y1. by subst. | ||
| 414 | - exists z. split. | ||
| 415 | + by apply rtc_once. | ||
| 416 | + by apply rc_once. | ||
| 417 | Qed. | ||
| 418 | |||
| 419 | Lemma step_semi_confluent : semi_confluent step. | ||
| 420 | Proof. apply strong__semi_confluence. apply step_strongly_confluent. Qed. | ||
| 421 | |||
| 422 | Lemma step_cr : cr step. | ||
| 423 | Proof. apply semi_confluence__cr. apply step_semi_confluent. Qed. | ||
| 424 | |||
| 425 | Theorem step_confluent : confluent step. | ||
| 426 | Proof. apply confluent_alt. apply step_cr. Qed. | ||