aboutsummaryrefslogtreecommitdiffstats
path: root/relations.v
diff options
context:
space:
mode:
authorLibravatar Rutger Broekhoff2024-06-27 01:56:05 +0200
committerLibravatar Rutger Broekhoff2024-06-27 01:56:05 +0200
commit4598e77a9406d8040357c943a05dd1ab939932db (patch)
treee990beaa94803da3585547f22e186e8819f6a887 /relations.v
parent73df1945b31c0beee88cf4476df4ccd09d31403b (diff)
downloadmininix-formalization-4598e77a9406d8040357c943a05dd1ab939932db.tar.gz
mininix-formalization-4598e77a9406d8040357c943a05dd1ab939932db.zip
Tidy up project
* Write README.md * Add LICENSE * Clean up some theorems and proofs
Diffstat (limited to 'relations.v')
-rw-r--r--relations.v19
1 files changed, 12 insertions, 7 deletions
diff --git a/relations.v b/relations.v
index 9bfb7e6..7523c96 100644
--- a/relations.v
+++ b/relations.v
@@ -1,10 +1,14 @@
1(** This files contains definitions for determinism, strong confluence,
2 * semi-confluence, the Church-Rosser property and proofs about relationships
3 * of these properties. We use definitions and proofs from 'Term Rewriting and
4 * All That' by Franz Baader and Tobias Nipkow (1998). *)
1From stdpp Require Import relations. 5From stdpp Require Import relations.
2 6
3Definition deterministic {A} (R : relation A) := 7Definition deterministic {A} (R : relation A) :=
4 ∀ x y1 y2, R x y1 → R x y2 → y1 = y2. 8 ∀ x y1 y2, R x y1 → R x y2 → y1 = y2.
5 9
6(* As in Baader and Nipkow (1998): *) 10(** Baader and Nipkow (1998), p. 9:
7(* y is a normal form of x if x -->* y and y is in normal form. *) 11 * y is a normal form of x if x -->* y and y is in normal form. *)
8Definition is_nf_of `(R : relation A) x y := (rtc R x y) ∧ nf R y. 12Definition is_nf_of `(R : relation A) x y := (rtc R x y) ∧ nf R y.
9 13
10(** The reflexive closure. *) 14(** The reflexive closure. *)
@@ -14,15 +18,15 @@ Inductive rc {A} (R : relation A) : relation A :=
14 18
15Hint Constructors rc : core. 19Hint Constructors rc : core.
16 20
17(* Baader and Nipkow, Definition 2.7.3. *) 21(** Baader and Nipkow, Definition 2.7.3. *)
18Definition strongly_confluent {A} (R : relation A) := 22Definition strongly_confluent {A} (R : relation A) :=
19 ∀ x y1 y2, R x y1 → R x y2 → ∃ z, rtc R y1 z ∧ rc R y2 z. 23 ∀ x y1 y2, R x y1 → R x y2 → ∃ z, rtc R y1 z ∧ rc R y2 z.
20 24
21(* Baader and Nipkow, Definition 2.1.4. *) 25(** Baader and Nipkow, Definition 2.1.4. *)
22Definition semi_confluent {A} (R : relation A) := 26Definition semi_confluent {A} (R : relation A) :=
23 ∀ x y1 y2, R x y1 → rtc R x y2 → ∃ z, rtc R y1 z ∧ rtc R y2 z. 27 ∀ x y1 y2, R x y1 → rtc R x y2 → ∃ z, rtc R y1 z ∧ rtc R y2 z.
24 28
25(* Lemma 2.7.4 from Baader and Nipkow *) 29(** Lemma 2.7.4 from Baader and Nipkow *)
26Lemma strong__semi_confluence {A} (R : relation A) : 30Lemma strong__semi_confluence {A} (R : relation A) :
27 strongly_confluent R → semi_confluent R. 31 strongly_confluent R → semi_confluent R.
28Proof. 32Proof.
@@ -47,11 +51,12 @@ Proof.
47 by apply rtc_transitive with (y := y2). 51 by apply rtc_transitive with (y := y2).
48Qed. 52Qed.
49 53
50(* Copied from stdpp, originates from Baader and Nipkow (1998) *) 54(** Baader and Nipkow, Definition 2.1.3.
55 * Copied from stdpp v1.10.0 (confluent_alt). *)
51Definition cr {A} (R : relation A) := 56Definition cr {A} (R : relation A) :=
52 ∀ x y, rtsc R x y → ∃ z, rtc R x z ∧ rtc R y z. 57 ∀ x y, rtsc R x y → ∃ z, rtc R x z ∧ rtc R y z.
53 58
54(* Part of Lemma 2.1.5 from Baader and Nipkow (1998) *) 59(** Part of Lemma 2.1.5 from Baader and Nipkow (1998) *)
55Lemma semi_confluence__cr {A} (R : relation A) : 60Lemma semi_confluence__cr {A} (R : relation A) :
56 semi_confluent R → cr R. 61 semi_confluent R → cr R.
57Proof. 62Proof.