Nice naive infrastructure for setoid rewriting.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "logic/cprop_connectives.ma".
+
+record equivalence_relation (A:Type) : Type ≝
+ { eq_rel:2> A → A → CProp;
+ refl: reflexive ? eq_rel;
+ sym: symmetric ? eq_rel;
+ trans: transitive ? eq_rel
+ }.
+
+record setoid : Type ≝
+ { carr:> Type;
+ eq: equivalence_relation carr
+ }.
+
+interpretation "setoid eq" 'eq x y = (eq_rel _ (eq _) x y).
+notation ".= r" with precedence 50 for @{trans ????? $r}.
+interpretation "setoid symmetry" 'invert r = (sym ____ r).
+
+record binary_morphism (A,B,C: setoid) : Type ≝
+ { fun:2> A → B → C;
+ prop: ∀a,a',b,b'. eq ? a a' → eq ? b b' → eq ? (fun a b) (fun a' b')
+ }.
+
+notation "# r" with precedence 60 for @{prop ???????? (refl ???) $r}.
+notation "r #" with precedence 60 for @{prop ???????? $r (refl ???)}.
+
+definition CPROP: setoid.
+ constructor 1;
+ [ apply CProp
+ | constructor 1;
+ [ apply Iff
+ | intros 1; split; intro; assumption
+ | intros 3; cases H; split; assumption
+ | intros 5; cases H; cases H1; split; intro;
+ [ apply (H4 (H2 H6)) | apply (H3 (H5 H6))]]]
+qed.
+
+definition eq_morphism: ∀S:setoid. binary_morphism S S CPROP.
+ intro;
+ constructor 1;
+ [ apply (eq_rel ? (eq S))
+ | intros; split; intro;
+ [ apply (.= H \sup -1);
+ apply (.= H2);
+ assumption
+ | apply (.= H);
+ apply (.= H2);
+ apply (H1 \sup -1)]]
+qed.
+
+record category : Type ≝
+ { objs:> Type;
+ arrows: objs → objs → setoid;
+ id: ∀o:objs. arrows o o;
+ comp: ∀o1,o2,o3. binary_morphism (arrows o1 o2) (arrows o2 o3) (arrows o1 o3);
+ comp_assoc: ∀o1,o2,o3,o4. ∀a12,a23,a34.
+ comp o1 o3 o4 (comp o1 o2 o3 a12 a23) a34 = comp o1 o2 o4 a12 (comp o2 o3 o4 a23 a34);
+ id_neutral_left: ∀o1,o2. ∀a: arrows o1 o2. comp ??? (id o1) a = a;
+ id_neutral_right: ∀o1,o2. ∀a: arrows o1 o2. comp ??? a (id o2) = a
+ }.
+
+interpretation "category composition" 'compose x y = (fun ___ (comp ____) x y).
+notation "'ASSOC'" with precedence 90 for @{comp_assoc ????????}.
\ No newline at end of file
(**************************************************************************)
include "logic/cprop_connectives.ma".
+include "datatypes/categories.ma".
record powerset (A: Type) : Type ≝ { char: A → CProp }.
interpretation "union" 'union U V = (union _ U V).
-record ssigma (A:Type) (S: powerset A) : Type ≝
- { witness:> A;
- proof:> witness ∈ S
- }.
-
-coercion ssigma.
-
-record binary_relation (A,B: Type) (U: Ω \sup A) (V: Ω \sup B) : Type ≝
- { satisfy:2> U → V → CProp }.
-
-(*notation < "hvbox (x (\circ term 19 r \frac \nbsp \circ) y)" with precedence 45 for @{'satisfy $r $x $y}.*)
-notation < "hvbox (x \nbsp \natur term 90 r \nbsp y)" with precedence 45 for @{'satisfy $r $x $y}.
-notation > "hvbox (x \natur term 90 r y)" with precedence 45 for @{'satisfy $r $x $y}.
-interpretation "relation applied" 'satisfy r x y = (satisfy ____ r x y).
-
-definition composition:
- ∀A,B,C.∀U1: Ω \sup A.∀U2: Ω \sup B.∀U3: Ω \sup C.
- binary_relation ?? U1 U2 → binary_relation ?? U2 U3 →
- binary_relation ?? U1 U3.
- intros (A B C U1 U2 U3 R12 R23);
- constructor 1;
- intros (s1 s3);
- apply (∃s2. s1 ♮R12 s2 ∧ s2 ♮R23 s3);
-qed.
-
-interpretation "binary relation composition" 'compose x y = (composition ______ x y).
-
-definition equal_relations ≝
- λA,B,U,V.λr,r': binary_relation A B U V.
- ∀x,y. r x y ↔ r' x y.
-
-interpretation "equal relation" 'eq x y = (equal_relations ____ x y).
-
-lemma refl_equal_relations: ∀A,B,U,V. reflexive ? (equal_relations A B U V).
- intros 5; intros 2; split; intro; assumption.
-qed.
-
-lemma sym_equal_relations: ∀A,B,U,V. symmetric ? (equal_relations A B U V).
- intros 7; intros 2; split; intro;
- [ apply (fi ?? (H ??)) | apply (if ?? (H ??))] assumption.
-qed.
-
-lemma trans_equal_relations: ∀A,B,U,V. transitive ? (equal_relations A B U V).
- intros 9; intros 2; split; intro;
- [ apply (if ?? (H1 ??)) | apply (fi ?? (H ??)) ]
- [ apply (if ?? (H ??)) | apply (fi ?? (H1 ??)) ]
- assumption.
-qed.
-
-lemma equal_morphism:
- ∀A,B,U,V.∀r1,r1',r2,r2':binary_relation A B U V.
- r1' = r1 → r2 = r2' → r1 = r2 → r1' = r2'.
- intros 13;
- split; intro;
- [ apply (if ?? (H1 ??));
- apply (if ?? (H2 ??));
- apply (if ?? (H ??));
- assumption
- | apply (fi ?? (H ??));
- apply (fi ?? (H2 ??));
- apply (fi ?? (H1 ??));
- assumption
- ]
-qed.
-
-lemma associative_composition:
- ∀A,B,C,D.∀U1,U2,U3,U4.
- ∀r1:binary_relation A B U1 U2.
- ∀r2:binary_relation B C U2 U3.
- ∀r3:binary_relation C D U3 U4.
- (r1 ∘ r2) ∘ r3 = r1 ∘ (r2 ∘ r3).
- intros 13;
- split; intro;
- cases H; clear H; cases H1; clear H1;
- [cases H; clear H | cases H2; clear H2]
- cases H1; clear H1;
- exists; try assumption;
- split; try assumption;
- exists; try assumption;
- split; assumption.
-qed.
-
-lemma composition_morphism:
- ∀A,B,C.∀U1,U2,U3.
- ∀r1,r1':binary_relation A B U1 U2.
- ∀r2,r2':binary_relation B C U2 U3.
- r1 = r1' → r2 = r2' → r1 ∘ r2 = r1' ∘ r2'.
- intros 14; split; intro;
- cases H2; clear H2; cases H3; clear H3;
- [ lapply (if ?? (H x w) H2) | lapply (fi ?? (H x w) H2) ]
- [ lapply (if ?? (H1 w y) H4)| lapply (fi ?? (H1 w y) H4) ]
- exists; try assumption;
- split; assumption.
-qed.
-
include "logic/equality.ma".
definition singleton ≝ λA:Type.λa:A.{b | a=b}.
-interpretation "singleton" 'singl a = (singleton _ a).
+interpretation "singleton" 'singl a = (singleton _ a).
\ No newline at end of file
(* *)
(**************************************************************************)
-include "datatypes/subsets.ma".
-include "logic/cprop_connectives.ma".
-include "formal_topology/categories.ma".
+include "formal_topology/relations.ma".
+include "datatypes/categories.ma".
record basic_pair: Type ≝
- { carr1: Type;
- carr2: Type;
- concr: Ω \sup carr1;
- form: Ω \sup carr2;
- rel: binary_relation ?? concr form
+ { concr: REL;
+ form: REL;
+ rel: arrows ? concr form
}.
notation "x ⊩ y" with precedence 45 for @{'Vdash2 $x $y}.
interpretation "basic pair relation" 'Vdash2 x y = (rel _ x y).
interpretation "basic pair relation (non applied)" 'Vdash = (rel _).
-alias symbol "eq" = "equal relation".
-alias symbol "compose" = "binary relation composition".
record relation_pair (BP1,BP2: basic_pair): Type ≝
- { concr_rel: binary_relation ?? (concr BP1) (concr BP2);
- form_rel: binary_relation ?? (form BP1) (form BP2);
+ { concr_rel: arrows ? (concr BP1) (concr BP2);
+ form_rel: arrows ? (form BP1) (form BP2);
commute: concr_rel ∘ ⊩ = ⊩ ∘ form_rel
}.
interpretation "concrete relation" 'concr_rel r = (concr_rel __ r).
interpretation "formal relation" 'form_rel r = (form_rel __ r).
-
definition relation_pair_equality:
∀o1,o2. equivalence_relation (relation_pair o1 o2).
intros;
[ apply (λr,r'. r \sub\c ∘ ⊩ = r' \sub\c ∘ ⊩);
| simplify;
intros;
- apply refl_equal_relations;
+ apply refl;
| simplify;
- intros;
- apply sym_equal_relations;
- assumption
+ intros 2;
+ apply sym;
| simplify;
- intros;
- apply (trans_equal_relations ??????? H);
- assumption
+ intros 3;
+ apply trans;
]
qed.
]
qed.
-definition eq' ≝
- λo1,o2.λr,r':relation_pair o1 o2.⊩ ∘ r \sub\f = ⊩ ∘ r' \sub\f.
-
-alias symbol "eq" = "setoid eq".
-lemma eq_to_eq': ∀o1,o2.∀r,r': relation_pair_setoid o1 o2. r=r' → eq' ?? r r'.
+lemma eq_to_eq': ∀o1,o2.∀r,r': relation_pair_setoid o1 o2. r=r' → ⊩ \circ r \sub\f = ⊩ \circ r'\sub\f.
intros 7 (o1 o2 r r' H c1 f2);
split; intro;
[ lapply (fi ?? (commute ?? r c1 f2) H1) as H2;
apply (if ?? (commute ?? r c1 f2) H3);
]
qed.
-
definition id: ∀o:basic_pair. relation_pair o o.
intro;
constructor 1;
- [1,2: constructor 1;
- intros;
- apply (s=s1)
- | simplify; intros;
- split;
- intro;
- cases H;
- cases H1; clear H H1;
- [ exists [ apply y ]
- split
- [ rewrite > H2; assumption
- | reflexivity ]
- | exists [ apply x ]
- split
- [2: rewrite < H3; assumption
- | reflexivity ]]]
+ [1,2: apply id;
+ | lapply (id_neutral_left ? (concr o) ? (⊩)) as H;
+ lapply (id_neutral_right ?? (form o) (⊩)) as H1;
+ apply (.= H);
+ apply (H1 \sup -1);]
qed.
definition relation_pair_composition:
| apply (r \sub\f ∘ r1 \sub\f)
| lapply (commute ?? r) as H;
lapply (commute ?? r1) as H1;
- apply (equal_morphism ???? (r\sub\c ∘ (r1\sub\c ∘ ⊩)) ? ((⊩ ∘ r\sub\f) ∘ r1\sub\f));
- [1,2: apply associative_composition]
- apply (equal_morphism ???? (r\sub\c ∘ (⊩ ∘ r1\sub\f)) ? ((r\sub\c ∘ ⊩) ∘ r1\sub\f));
- [1,2: apply composition_morphism; first [assumption | apply refl_equal_relations]
- | apply sym_equal_relations;
- apply associative_composition
- ]]
+ apply (.= ASSOC);
+ apply (.= #H1);
+ apply (.= ASSOC\sup -1);
+ apply (.= H#);
+ apply comp_assoc]
| intros;
- alias symbol "eq" = "equal relation".
- change with (a\sub\c ∘ b\sub\c ∘ ⊩ = a'\sub\c ∘ b'\sub\c ∘ ⊩);
- apply (equal_morphism ???? (a\sub\c ∘ (b\sub\c ∘ ⊩)) ? (a'\sub\c ∘ (b'\sub\c ∘ ⊩)));
- [ apply associative_composition
- | apply sym_equal_relations; apply associative_composition]
- apply (equal_morphism ???? (a\sub\c ∘ (b'\sub\c ∘ ⊩)) ? (a' \sub \c∘(b' \sub \c∘⊩)));
- [2: apply refl_equal_relations;
- |1: apply composition_morphism;
- [ apply refl_equal_relations
- | assumption]]
- apply (equal_morphism ???? (a\sub\c ∘ (⊩ ∘ b'\sub\f)) ? (a'\sub\c ∘ (⊩ ∘ b'\sub\f)));
- [1,2: apply composition_morphism;
- [1,3: apply refl_equal_relations
- | apply (commute ?? b');
- | apply sym_equal_relations; apply (commute ?? b');]]
- apply (equal_morphism ???? ((a\sub\c ∘ ⊩) ∘ b'\sub\f) ? ((a'\sub\c ∘ ⊩) ∘ b'\sub\f));
- [2: apply associative_composition
- |1: apply sym_equal_relations; apply associative_composition]
- apply composition_morphism;
- [ assumption
- | apply refl_equal_relations]]
+ change with (a\sub\c ∘ b\sub\c ∘ ⊩ = a'\sub\c ∘ b'\sub\c ∘ ⊩);
+ change in H with (a \sub\c ∘ ⊩ = a' \sub\c ∘ ⊩);
+ change in H1 with (b \sub\c ∘ ⊩ = b' \sub\c ∘ ⊩);
+ apply (.= ASSOC);
+ apply (.= #H1);
+ apply (.= #(commute ?? b'));
+ apply (.= ASSOC \sup -1);
+ apply (.= H#);
+ apply (.= ASSOC);
+ apply (.= #(commute ?? b')\sup -1);
+ apply (ASSOC \sup -1)]
qed.
-
+
definition BP: category.
constructor 1;
[ apply basic_pair
| intros;
change with (a12\sub\c ∘ a23\sub\c ∘ a34\sub\c ∘ ⊩ =
(a12\sub\c ∘ (a23\sub\c ∘ a34\sub\c) ∘ ⊩));
- apply composition_morphism;
- [2: apply refl_equal_relations]
- apply associative_composition
+ apply (ASSOC#);
| intros;
change with ((id o1)\sub\c ∘ a\sub\c ∘ ⊩ = a\sub\c ∘ ⊩);
- apply composition_morphism;
- [2: apply refl_equal_relations]
- intros 2; unfold id; simplify;
- split; intro;
- [ cases H; cases H1; rewrite > H2; assumption
- | exists; [assumption] split; [reflexivity| assumption]]
+ apply ((id_neutral_left ????)#);
| intros;
change with (a\sub\c ∘ (id o2)\sub\c ∘ ⊩ = a\sub\c ∘ ⊩);
- apply composition_morphism;
- [2: apply refl_equal_relations]
- intros 2; unfold id; simplify;
- split; intro;
- [ cases H; cases H1; rewrite < H3; assumption
- | exists; [assumption] split; [assumption|reflexivity]]
+ apply ((id_neutral_right ????)#);
]
-qed.
+qed.
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "logic/cprop_connectives.ma".
-
-record equivalence_relation (A:Type) : Type ≝
- { eq_rel:2> A → A → CProp;
- refl: reflexive ? eq_rel;
- sym: symmetric ? eq_rel;
- trans: transitive ? eq_rel
- }.
-
-record setoid : Type ≝
- { carr:> Type;
- eq: equivalence_relation carr
- }.
-
-interpretation "setoid eq" 'eq x y = (eq_rel _ (eq _) x y).
-
-record binary_morphism (A,B,C: setoid) : Type ≝
- { fun:2> A → B → C;
- prop: ∀a,a',b,b'. eq ? a a' → eq ? b b' → eq ? (fun a b) (fun a' b')
- }.
-
-record category : Type ≝
- { objs: Type;
- arrows: objs → objs → setoid;
- id: ∀o:objs. arrows o o;
- comp: ∀o1,o2,o3. binary_morphism (arrows o1 o2) (arrows o2 o3) (arrows o1 o3);
- comp_assoc: ∀o1,o2,o3,o4. ∀a12,a23,a34.
- comp o1 o3 o4 (comp o1 o2 o3 a12 a23) a34 = comp o1 o2 o4 a12 (comp o2 o3 o4 a23 a34);
- id_neutral_left: ∀o1,o2. ∀a: arrows o1 o2. comp ??? (id o1) a = a;
- id_neutral_right: ∀o1,o2. ∀a: arrows o1 o2. comp ??? a (id o2) = a
- }.
-
-interpretation "category composition" 'compose x y = (comp ____ x y).
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "datatypes/subsets.ma".
+
+record ssigma (A:Type) (S: powerset A) : Type ≝
+ { witness:> A;
+ proof:> witness ∈ S
+ }.
+
+coercion ssigma.
+
+record binary_relation (A,B: Type) (U: Ω \sup A) (V: Ω \sup B) : Type ≝
+ { satisfy:2> U → V → CProp }.
+
+notation < "hvbox (x \nbsp \natur term 90 r \nbsp y)" with precedence 45 for @{'satisfy $r $x $y}.
+notation > "hvbox (x \natur term 90 r y)" with precedence 45 for @{'satisfy $r $x $y}.
+interpretation "relation applied" 'satisfy r x y = (satisfy ____ r x y).
+
+definition composition:
+ ∀A,B,C.∀U1: Ω \sup A.∀U2: Ω \sup B.∀U3: Ω \sup C.
+ binary_relation ?? U1 U2 → binary_relation ?? U2 U3 →
+ binary_relation ?? U1 U3.
+ intros (A B C U1 U2 U3 R12 R23);
+ constructor 1;
+ intros (s1 s3);
+ apply (∃s2. s1 ♮R12 s2 ∧ s2 ♮R23 s3);
+qed.
+
+interpretation "binary relation composition" 'compose x y = (composition ______ x y).
+
+definition equal_relations ≝
+ λA,B,U,V.λr,r': binary_relation A B U V.
+ ∀x,y. r x y ↔ r' x y.
+
+interpretation "equal relation" 'eq x y = (equal_relations ____ x y).
+
+lemma refl_equal_relations: ∀A,B,U,V. reflexive ? (equal_relations A B U V).
+ intros 5; intros 2; split; intro; assumption.
+qed.
+
+lemma sym_equal_relations: ∀A,B,U,V. symmetric ? (equal_relations A B U V).
+ intros 7; intros 2; split; intro;
+ [ apply (fi ?? (H ??)) | apply (if ?? (H ??))] assumption.
+qed.
+
+lemma trans_equal_relations: ∀A,B,U,V. transitive ? (equal_relations A B U V).
+ intros 9; intros 2; split; intro;
+ [ apply (if ?? (H1 ??)) | apply (fi ?? (H ??)) ]
+ [ apply (if ?? (H ??)) | apply (fi ?? (H1 ??)) ]
+ assumption.
+qed.
+
+lemma associative_composition:
+ ∀A,B,C,D.∀U1,U2,U3,U4.
+ ∀r1:binary_relation A B U1 U2.
+ ∀r2:binary_relation B C U2 U3.
+ ∀r3:binary_relation C D U3 U4.
+ (r1 ∘ r2) ∘ r3 = r1 ∘ (r2 ∘ r3).
+ intros 13;
+ split; intro;
+ cases H; clear H; cases H1; clear H1;
+ [cases H; clear H | cases H2; clear H2]
+ cases H1; clear H1;
+ exists; try assumption;
+ split; try assumption;
+ exists; try assumption;
+ split; assumption.
+qed.
+
+lemma composition_morphism:
+ ∀A,B,C.∀U1,U2,U3.
+ ∀r1,r1':binary_relation A B U1 U2.
+ ∀r2,r2':binary_relation B C U2 U3.
+ r1 = r1' → r2 = r2' → r1 ∘ r2 = r1' ∘ r2'.
+ intros 14; split; intro;
+ cases H2; clear H2; cases H3; clear H3;
+ [ lapply (if ?? (H x w) H2) | lapply (fi ?? (H x w) H2) ]
+ [ lapply (if ?? (H1 w y) H4)| lapply (fi ?? (H1 w y) H4) ]
+ exists; try assumption;
+ split; assumption.
+qed.
+
+definition binary_relation_setoid: ∀A,B. Ω \sup A → Ω \sup B → setoid.
+ intros (A B U V);
+ constructor 1;
+ [ apply (binary_relation ?? U V)
+ | constructor 1;
+ [ apply equal_relations
+ | apply refl_equal_relations
+ | apply sym_equal_relations
+ | apply trans_equal_relations
+ ]]
+qed.
+
+record sigma (A:Type) (P: A → Type) : Type ≝
+ { s_witness:> A;
+ s_proof:> P s_witness
+ }.
+
+interpretation "sigma" 'sigma \eta.x = (sigma _ x).
+
+definition REL: category.
+ constructor 1;
+ [ apply (ΣA:Type.Ω \sup A)
+ | intros; apply (binary_relation_setoid ?? (s_proof ?? s) (s_proof ?? s1))
+ | intros; constructor 1; intros; apply (s=s1)
+ | intros; constructor 1;
+ [ apply composition
+ | apply composition_morphism
+ ]
+ | intros; unfold mk_binary_morphism; simplify;
+ apply associative_composition
+ |6,7: intros 5; simplify; split; intro;
+ [1,3: cases H; clear H; cases H1; clear H1;
+ [ rewrite > H | rewrite < H2 ]
+ assumption
+ |*: exists; try assumption; split; first [ reflexivity | assumption ]]]
+qed.
\ No newline at end of file