33From HB Require Import structures.
44From Coq Require Import ssreflect ssrfun.
55
6- HB.mixin Record isQuiver Obj := { hom : Obj -> Obj -> Type }.
6+ HB.mixin Record isQuiver ( Obj: Type) : Type := { hom : Obj -> Obj -> Type }.
77
8- HB.structure Definition Quiver := { Obj of isQuiver Obj }.
8+ HB.structure Definition Quiver : Type : = { Obj of isQuiver Obj }.
99
10- HB.mixin Record isMon A := {
10+ HB.mixin Record isMon (A: Type ) : Type := {
1111 zero : A;
1212 add : A -> A -> A;
1313 addrA : associative add;
@@ -16,20 +16,94 @@ HB.mixin Record isMon A := {
1616 }.
1717
1818HB.structure
19- Definition Monoid := { A of isMon A }.
19+ Definition Monoid : Type : = { A of isMon A }.
2020
21- (**)
22- HB.mixin Record hom_isMonT T of Quiver T :=
21+ HB.mixin Record hom_isMon T of Quiver T :=
2322 { private : forall A B, isMon (@hom T A B) }.
2423
2524HB.structure
26- Definition Monoid_enriched_quiverT :=
27- { Obj of isQuiver Obj & hom_isMonT Obj }.
25+ Definition Monoid_enriched_quiver :=
26+ { Obj of isQuiver Obj & hom_isMon Obj }.
27+
28+ (* unique projection from the axiom of Monoid_enriched_quiver *)
29+ HB.instance Definition _ (T : Monoid_enriched_quiver.type) (A B : T) : isMon (@hom T A B) :=
30+ @private T A B.
31+
32+ HB.instance Definition _ := isQuiver.Build Type (fun A B => A -> B).
33+
34+
35+ (******** *)
36+
37+ HB.mixin Record hom_isMonX T of Quiver T : Type :=
38+ { private : forall A B, isMon (@hom T A B) }.
39+
40+ HB.structure
41+ Definition Monoid_enriched_quiverX :=
42+ { Obj of isQuiver Obj & hom_isMonX Obj }.
43+
44+ Record isQuiverS (Obj: Type) : Type := { homS : Obj -> Obj -> Type }.
45+
46+ Structure QuiverS := { ObjS: Type; AxS: isQuiverS ObjS }.
47+
48+ Definition hom_isMon_type T (X: isQuiverS T) (A B: T) : Type :=
49+ isMon (@homS T X A B).
50+
51+ Record hom_isMonQ T (X: isQuiverS T) : Type :=
52+ hiMQ { privateQ : forall (A B: T), hom_isMon_type T X A B }.
53+
54+ Definition my_hom_isMonQ T (X: isQuiverS T) (F: forall A B, hom_isMon_type T X A B) :
55+ hom_isMonQ T X := hiMQ T X F.
56+
57+ Record Monoid_enriched_quiverQ := { ObjQ: Type; iQQ: isQuiverS ObjQ; hsM: hom_isMonQ ObjQ iQQ }.
58+
59+ Record hom_wrapper T (X: isQuiverS T) (Str: Type -> Type ) : Type :=
60+ { privateW : forall (A B: T), Str (@homS T X A B) }.
2861
29- (**)
62+ Record hom_wrapperA T (Qv: Type -> Type ) (hm: Qv T -> T -> T -> Type ) (Str: Type -> Type ) (x: Qv T) : Type :=
63+ { privateWA : forall (A B: T), Str (hm x A B) }.
3064
65+ Definition my_quiver (T: Type) : isQuiverS T.
66+ Admitted .
3167
32- Fail HB.structure
68+ Lemma my_quiver_mon (T: Type) : forall (A B: T), isMon (@homS T (my_quiver T) A B).
69+ Admitted .
70+
71+ Definition my_hom_isMon (T: Type ) : hom_isMonQ T (my_quiver T) :=
72+ my_hom_isMonQ T (my_quiver T) (my_quiver_mon T).
73+
74+ Definition Mixin : Type := Type -> Type.
75+
76+ (* write two versions of Monoid_enriched_quiver: one using hom_isMon
77+ (a mixin, hence a record), the other one using hom_isMon_type as naked
78+ field type. The former corresponds to the wrapped version, the latter
79+ to the intuitive, wrapper-less version. Now the latter should agree
80+ with the former... (broadly corresponding to 2?) *)
81+
82+
83+
84+ Lemma quiver_ok (T: Type ) (Str: Type -> Type ) :
85+ forall (A B: T), @homS (my_quiver T) A B
86+
87+
88+ Class wrapper (T: Type) () (P: T -> T -> Prop) { prop: forall A B: T, P A B }.
89+
90+ Definition wrapped_hom (T: Type ) (F: isMon (@hom T A B)) := wrapper T (isMon (@hom T))
91+
92+
93+ Elpi Accumulate lp:{{
94+
95+ pred wrapper-mixin o:mixinname, o:gref, o:mixinname.
96+
97+ }}.
98+
99+
100+
101+ HB.instance Definition homTypeMon (A B : Quiver.type) := isMon.Build (hom A B) (* ... *) .
102+
103+ (******** *)
104+
105+
106+ HB.structure
33107 Definition Monoid_enriched_quiver :=
34108 { Obj of isQuiver Obj &
35109 (forall A B : Obj, isMon (@hom (Quiver.clone Obj _) A B)) }.
0 commit comments