Skip to content

Commit 1cac896

Browse files
committed
1 parent f9712ba commit 1cac896

21 files changed

+181
-180
lines changed

implementations/fast_rationals.v

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,9 @@
11
Require
2+
MathClasses.theory.rings
23
MathClasses.theory.shiftl MathClasses.theory.int_pow.
34
Require Import
45
Coq.QArith.QArith Bignums.BigQ.BigQ
5-
MathClasses.interfaces.abstract_algebra
6+
MathClasses.interfaces.orders MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals
67
MathClasses.interfaces.integers MathClasses.interfaces.rationals MathClasses.interfaces.additional_operations
78
MathClasses.implementations.fast_naturals MathClasses.implementations.fast_integers MathClasses.implementations.field_of_fractions MathClasses.implementations.stdlib_rationals.
89
Require Export

interfaces/abstract_algebra.v

Lines changed: 71 additions & 71 deletions
Original file line numberDiff line numberDiff line change
@@ -6,9 +6,9 @@ For various structures we omit declaration of substructures. For example, if we
66
say:
77
88
Class Setoid_Morphism :=
9-
{ setoidmor_a :> Setoid A
10-
; setoidmor_b :> Setoid B
11-
; sm_proper :> Proper ((=) ==> (=)) f }.
9+
{ setoidmor_a :: Setoid A
10+
; setoidmor_b :: Setoid B
11+
; sm_proper :: Proper ((=) ==> (=)) f }.
1212
1313
then each time a Setoid instance is required, Coq will try to prove that a
1414
Setoid_Morphism exists. This obviously results in an enormous blow-up of the
@@ -20,14 +20,14 @@ Setoid_Morphism as a substructure, setoid rewriting will become horribly slow.
2020
*)
2121

2222
(* An unbundled variant of the former CoRN RSetoid *)
23-
Class Setoid A {Ae : Equiv A} : Prop := setoid_eq :> Equivalence (@equiv A Ae).
23+
Class Setoid A {Ae : Equiv A} : Prop := setoid_eq :: Equivalence (@equiv A Ae).
2424

2525
(* An unbundled variant of the former CoRN CSetoid. We do not
2626
include a proof that A is a Setoid because it can be derived. *)
2727
Class StrongSetoid A {Ae : Equiv A} `{Aap : Apart A} : Prop :=
28-
{ strong_setoid_irreflexive :> Irreflexive (≶)
29-
; strong_setoid_symmetric :> Symmetric (≶)
30-
; strong_setoid_cotrans :> CoTransitive (≶)
28+
{ strong_setoid_irreflexive :: Irreflexive (≶)
29+
; strong_setoid_symmetric :: Symmetric (≶)
30+
; strong_setoid_cotrans :: CoTransitive (≶)
3131
; tight_apart : ∀ x y, ¬x ≶ y ↔ x = y }.
3232

3333
Arguments tight_apart {A Ae Aap StrongSetoid} _ _.
@@ -38,7 +38,7 @@ Section setoid_morphisms.
3838
Class Setoid_Morphism :=
3939
{ setoidmor_a : Setoid A
4040
; setoidmor_b : Setoid B
41-
; sm_proper :> Proper ((=) ==> (=)) f }.
41+
; sm_proper :: Proper ((=) ==> (=)) f }.
4242

4343
Class StrongSetoid_Morphism :=
4444
{ strong_setoidmor_a : StrongSetoid A
@@ -70,80 +70,80 @@ Section upper_classes.
7070
Context A {Ae : Equiv A}.
7171

7272
Class SemiGroup {Aop: SgOp A} : Prop :=
73-
{ sg_setoid :> Setoid A
74-
; sg_ass :> Associative (&)
75-
; sg_op_proper :> Proper ((=) ==> (=) ==> (=)) (&) }.
73+
{ sg_setoid :: Setoid A
74+
; sg_ass :: Associative (&)
75+
; sg_op_proper :: Proper ((=) ==> (=) ==> (=)) (&) }.
7676

7777
Class CommutativeSemiGroup {Aop : SgOp A} : Prop :=
78-
{ comsg_setoid :> @SemiGroup Aop
79-
; comsg_ass :> Commutative (&) }.
78+
{ comsg_setoid :: @SemiGroup Aop
79+
; comsg_ass :: Commutative (&) }.
8080

8181
Class SemiLattice {Aop : SgOp A} : Prop :=
82-
{ semilattice_sg :> @CommutativeSemiGroup Aop
83-
; semilattice_idempotent :> BinaryIdempotent (&)}.
82+
{ semilattice_sg :: @CommutativeSemiGroup Aop
83+
; semilattice_idempotent :: BinaryIdempotent (&)}.
8484

8585
Class Monoid {Aop : SgOp A} {Aunit : MonUnit A} : Prop :=
86-
{ monoid_semigroup :> SemiGroup
87-
; monoid_left_id :> LeftIdentity (&) mon_unit
88-
; monoid_right_id :> RightIdentity (&) mon_unit }.
86+
{ monoid_semigroup :: SemiGroup
87+
; monoid_left_id :: LeftIdentity (&) mon_unit
88+
; monoid_right_id :: RightIdentity (&) mon_unit }.
8989

9090
Class CommutativeMonoid {Aop Aunit} : Prop :=
91-
{ commonoid_mon :> @Monoid Aop Aunit
92-
; commonoid_commutative :> Commutative (&) }.
91+
{ commonoid_mon :: @Monoid Aop Aunit
92+
; commonoid_commutative :: Commutative (&) }.
9393

9494
Class Group {Aop Aunit} {Anegate : Negate A} : Prop :=
95-
{ group_monoid :> @Monoid Aop Aunit
96-
; negate_proper :> Setoid_Morphism (-)
97-
; negate_l :> LeftInverse (&) (-) mon_unit
98-
; negate_r :> RightInverse (&) (-) mon_unit }.
95+
{ group_monoid :: @Monoid Aop Aunit
96+
; negate_proper :: Setoid_Morphism (-)
97+
; negate_l :: LeftInverse (&) (-) mon_unit
98+
; negate_r :: RightInverse (&) (-) mon_unit }.
9999

100100
Class BoundedSemiLattice {Aop Aunit} : Prop :=
101-
{ bounded_semilattice_mon :> @CommutativeMonoid Aop Aunit
102-
; bounded_semilattice_idempotent :> BinaryIdempotent (&)}.
101+
{ bounded_semilattice_mon :: @CommutativeMonoid Aop Aunit
102+
; bounded_semilattice_idempotent :: BinaryIdempotent (&)}.
103103

104104
Class AbGroup {Aop Aunit Anegate} : Prop :=
105-
{ abgroup_group :> @Group Aop Aunit Anegate
106-
; abgroup_commutative :> Commutative (&) }.
105+
{ abgroup_group :: @Group Aop Aunit Anegate
106+
; abgroup_commutative :: Commutative (&) }.
107107

108108
Context {Aplus : Plus A} {Amult : Mult A} {Azero : Zero A} {Aone : One A}.
109109

110110
Class SemiRing : Prop :=
111-
{ semiplus_monoid :> @CommutativeMonoid plus_is_sg_op zero_is_mon_unit
112-
; semimult_monoid :> @CommutativeMonoid mult_is_sg_op one_is_mon_unit
113-
; semiring_distr :> LeftDistribute (.*.) (+)
114-
; semiring_left_absorb :> LeftAbsorb (.*.) 0 }.
111+
{ semiplus_monoid :: @CommutativeMonoid plus_is_sg_op zero_is_mon_unit
112+
; semimult_monoid :: @CommutativeMonoid mult_is_sg_op one_is_mon_unit
113+
; semiring_distr :: LeftDistribute (.*.) (+)
114+
; semiring_left_absorb :: LeftAbsorb (.*.) 0 }.
115115

116116
Context {Anegate : Negate A}.
117117

118118
Class Ring : Prop :=
119-
{ ring_group :> @AbGroup plus_is_sg_op zero_is_mon_unit _
120-
; ring_monoid :> @CommutativeMonoid mult_is_sg_op one_is_mon_unit
121-
; ring_dist :> LeftDistribute (.*.) (+) }.
119+
{ ring_group :: @AbGroup plus_is_sg_op zero_is_mon_unit _
120+
; ring_monoid :: @CommutativeMonoid mult_is_sg_op one_is_mon_unit
121+
; ring_dist :: LeftDistribute (.*.) (+) }.
122122

123123
(* For now, we follow CoRN/ring_theory's example in having Ring and SemiRing
124124
require commutative multiplication. *)
125125

126126
Class IntegralDomain : Prop :=
127127
{ intdom_ring : Ring
128128
; intdom_nontrivial : PropHolds (1 ≠ 0)
129-
; intdom_nozeroes :> NoZeroDivisors A }.
129+
; intdom_nozeroes :: NoZeroDivisors A }.
130130

131131
(* We do not include strong extensionality for (-) and (/) because it can de derived *)
132132
Class Field {Aap: Apart A} {Arecip: Recip A} : Prop :=
133-
{ field_ring :> Ring
134-
; field_strongsetoid :> StrongSetoid A
135-
; field_plus_ext :> StrongSetoid_BinaryMorphism (+)
136-
; field_mult_ext :> StrongSetoid_BinaryMorphism (.*.)
133+
{ field_ring :: Ring
134+
; field_strongsetoid :: StrongSetoid A
135+
; field_plus_ext :: StrongSetoid_BinaryMorphism (+)
136+
; field_mult_ext :: StrongSetoid_BinaryMorphism (.*.)
137137
; field_nontrivial : PropHolds (1 ≶ 0)
138-
; recip_proper :> Setoid_Morphism (//)
138+
; recip_proper :: Setoid_Morphism (//)
139139
; recip_inverse : ∀ x, `x // x = 1 }.
140140

141141
(* We let /0 = 0 so properties as Injective (/), f (/x) = / (f x), / /x = x, /x * /y = /(x * y)
142142
hold without any additional assumptions *)
143143
Class DecField {Adec_recip : DecRecip A} : Prop :=
144-
{ decfield_ring :> Ring
144+
{ decfield_ring :: Ring
145145
; decfield_nontrivial : PropHolds (1 ≠ 0)
146-
; dec_recip_proper :> Setoid_Morphism (/)
146+
; dec_recip_proper :: Setoid_Morphism (/)
147147
; dec_recip_0 : /0 = 0
148148
; dec_recip_inverse : ∀ x, x ≠ 0 → x / x = 1 }.
149149
End upper_classes.
@@ -159,7 +159,7 @@ Hint Extern 5 (PropHolds (1 ≠ 0)) => eapply @decfield_nontrivial : typeclass_i
159159
(*
160160
For a strange reason Ring instances of Integers are sometimes obtained by
161161
Integers -> IntegralDomain -> Ring and sometimes directly. Making this an
162-
instance with a low priority instead of using intdom_ring:> Ring forces Coq to
162+
instance with a low priority instead of using intdom_ring:: Ring forces Coq to
163163
take the right way
164164
*)
165165
#[global]
@@ -174,29 +174,29 @@ Section lattice.
174174
Context A `{Ae: Equiv A} {Ajoin: Join A} {Ameet: Meet A} `{Abottom : Bottom A}.
175175

176176
Class JoinSemiLattice : Prop :=
177-
join_semilattice :> @SemiLattice A Ae join_is_sg_op.
177+
join_semilattice :: @SemiLattice A Ae join_is_sg_op.
178178
Class BoundedJoinSemiLattice : Prop :=
179-
bounded_join_semilattice :> @BoundedSemiLattice A Ae join_is_sg_op bottom_is_mon_unit.
179+
bounded_join_semilattice :: @BoundedSemiLattice A Ae join_is_sg_op bottom_is_mon_unit.
180180
Class MeetSemiLattice : Prop :=
181-
meet_semilattice :> @SemiLattice A Ae meet_is_sg_op.
181+
meet_semilattice :: @SemiLattice A Ae meet_is_sg_op.
182182

183183
Class Lattice : Prop :=
184-
{ lattice_join :> JoinSemiLattice
185-
; lattice_meet :> MeetSemiLattice
186-
; join_meet_absorption :> Absorption (⊔) (⊓)
187-
; meet_join_absorption :> Absorption (⊓) (⊔)}.
184+
{ lattice_join :: JoinSemiLattice
185+
; lattice_meet :: MeetSemiLattice
186+
; join_meet_absorption :: Absorption (⊔) (⊓)
187+
; meet_join_absorption :: Absorption (⊓) (⊔)}.
188188

189189
Class DistributiveLattice : Prop :=
190-
{ distr_lattice_lattice :> Lattice
191-
; join_meet_distr_l :> LeftDistribute (⊔) (⊓) }.
190+
{ distr_lattice_lattice :: Lattice
191+
; join_meet_distr_l :: LeftDistribute (⊔) (⊓) }.
192192
End lattice.
193193

194194
Class Category O `{!Arrows O} `{∀ x y: O, Equiv (x ⟶ y)} `{!CatId O} `{!CatComp O}: Prop :=
195-
{ arrow_equiv :> ∀ x y, Setoid (x ⟶ y)
196-
; comp_proper :> ∀ x y z, Proper ((=) ==> (=) ==> (=)) (comp x y z)
197-
; comp_assoc :> ArrowsAssociative O
198-
; id_l :> ∀ x y, LeftIdentity (comp x y y) cat_id
199-
; id_r :> ∀ x y, RightIdentity (comp x x y) cat_id }.
195+
{ arrow_equiv :: ∀ x y, Setoid (x ⟶ y)
196+
; comp_proper :: ∀ x y z, Proper ((=) ==> (=) ==> (=)) (comp x y z)
197+
; comp_assoc :: ArrowsAssociative O
198+
; id_l :: ∀ x y, LeftIdentity (comp x y y) cat_id
199+
; id_r :: ∀ x y, RightIdentity (comp x x y) cat_id }.
200200
(* note: no equality on objects. *)
201201

202202
(* todo: use my comp everywhere *)
@@ -208,46 +208,46 @@ Section morphism_classes.
208208
Class SemiGroup_Morphism {Aop Bop} (f : A → B) :=
209209
{ sgmor_a : @SemiGroup A Ae Aop
210210
; sgmor_b : @SemiGroup B Be Bop
211-
; sgmor_setmor :> Setoid_Morphism f
211+
; sgmor_setmor :: Setoid_Morphism f
212212
; preserves_sg_op : ∀ x y, f (x & y) = f x & f y }.
213213

214214
Class JoinSemiLattice_Morphism {Ajoin Bjoin} (f : A → B) :=
215215
{ join_slmor_a : @JoinSemiLattice A Ae Ajoin
216216
; join_slmor_b : @JoinSemiLattice B Be Bjoin
217-
; join_slmor_sgmor :> @SemiGroup_Morphism join_is_sg_op join_is_sg_op f }.
217+
; join_slmor_sgmor :: @SemiGroup_Morphism join_is_sg_op join_is_sg_op f }.
218218

219219
Class MeetSemiLattice_Morphism {Ameet Bmeet} (f : A → B) :=
220220
{ meet_slmor_a : @MeetSemiLattice A Ae Ameet
221221
; meet_slmor_b : @MeetSemiLattice B Be Bmeet
222-
; meet_slmor_sgmor :> @SemiGroup_Morphism meet_is_sg_op meet_is_sg_op f }.
222+
; meet_slmor_sgmor :: @SemiGroup_Morphism meet_is_sg_op meet_is_sg_op f }.
223223

224224
Class Monoid_Morphism {Aunit Bunit Aop Bop} (f : A → B) :=
225225
{ monmor_a : @Monoid A Ae Aop Aunit
226226
; monmor_b : @Monoid B Be Bop Bunit
227-
; monmor_sgmor :> SemiGroup_Morphism f
227+
; monmor_sgmor :: SemiGroup_Morphism f
228228
; preserves_mon_unit : f mon_unit = mon_unit }.
229229

230230
Class BoundedJoinSemiLattice_Morphism {Abottom Bbottom Ajoin Bjoin} (f : A → B) :=
231231
{ bounded_join_slmor_a : @BoundedJoinSemiLattice A Ae Ajoin Abottom
232232
; bounded_join_slmor_b : @BoundedJoinSemiLattice B Be Bjoin Bbottom
233-
; bounded_join_slmor_monmor :> @Monoid_Morphism bottom_is_mon_unit bottom_is_mon_unit join_is_sg_op join_is_sg_op f }.
233+
; bounded_join_slmor_monmor :: @Monoid_Morphism bottom_is_mon_unit bottom_is_mon_unit join_is_sg_op join_is_sg_op f }.
234234

235235
Class SemiRing_Morphism {Aplus Amult Azero Aone Bplus Bmult Bzero Bone} (f : A → B) :=
236236
{ semiringmor_a : @SemiRing A Ae Aplus Amult Azero Aone
237237
; semiringmor_b : @SemiRing B Be Bplus Bmult Bzero Bone
238-
; semiringmor_plus_mor :> @Monoid_Morphism zero_is_mon_unit zero_is_mon_unit plus_is_sg_op plus_is_sg_op f
239-
; semiringmor_mult_mor :> @Monoid_Morphism one_is_mon_unit one_is_mon_unit mult_is_sg_op mult_is_sg_op f }.
238+
; semiringmor_plus_mor :: @Monoid_Morphism zero_is_mon_unit zero_is_mon_unit plus_is_sg_op plus_is_sg_op f
239+
; semiringmor_mult_mor :: @Monoid_Morphism one_is_mon_unit one_is_mon_unit mult_is_sg_op mult_is_sg_op f }.
240240

241241
Class Lattice_Morphism {Ajoin Ameet Bjoin Bmeet} (f : A → B) :=
242242
{ latticemor_a : @Lattice A Ae Ajoin Ameet
243243
; latticemor_b : @Lattice B Be Bjoin Bmeet
244-
; latticemor_join_mor :> JoinSemiLattice_Morphism f
245-
; latticemor_meet_mor :> MeetSemiLattice_Morphism f }.
244+
; latticemor_join_mor :: JoinSemiLattice_Morphism f
245+
; latticemor_meet_mor :: MeetSemiLattice_Morphism f }.
246246

247247
Context {Aap : Apart A} {Bap : Apart B}.
248248
Class StrongSemiRing_Morphism {Aplus Amult Azero Aone Bplus Bmult Bzero Bone} (f : A → B) :=
249-
{ strong_semiringmor_sr_mor :> @SemiRing_Morphism Aplus Amult Azero Aone Bplus Bmult Bzero Bone f
250-
; strong_semiringmor_strong_mor :> StrongSetoid_Morphism f }.
249+
{ strong_semiringmor_sr_mor :: @SemiRing_Morphism Aplus Amult Azero Aone Bplus Bmult Bzero Bone f
250+
; strong_semiringmor_strong_mor :: StrongSetoid_Morphism f }.
251251
End morphism_classes.
252252

253253
Section jections.
@@ -267,8 +267,8 @@ Section jections.
267267
; surjective_mor : Setoid_Morphism f }.
268268

269269
Class Bijective : Prop :=
270-
{ bijective_injective :> Injective
271-
; bijective_surjective :> Surjective }.
270+
{ bijective_injective :: Injective
271+
; bijective_surjective :: Surjective }.
272272
End jections.
273273

274274
#[global]

interfaces/additional_operations.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,7 @@ Instance: Params (@shiftl) 3 := {}.
3434

3535
Class ShiftLSpec A B (sl : ShiftL A B) `{Equiv A} `{Equiv B} `{One A} `{Plus A} `{Mult A} `{Zero B} `{One B} `{Plus B} := {
3636
shiftl_proper : Proper ((=) ==> (=) ==> (=)) (≪) ;
37-
shiftl_0 :> RightIdentity (≪) 0 ;
37+
shiftl_0 :: RightIdentity (≪) 0 ;
3838
shiftl_S : ∀ x n, x ≪ (1 + n) = 2 * x ≪ n
3939
}.
4040

@@ -71,7 +71,7 @@ Instance: Params (@shiftr) 3 := {}.
7171

7272
Class ShiftRSpec A B (sl : ShiftR A B) `{Equiv A} `{Equiv B} `{One A} `{Plus A} `{Mult A} `{Zero B} `{One B} `{Plus B} := {
7373
shiftr_proper : Proper ((=) ==> (=) ==> (=)) (≫) ;
74-
shiftr_0 :> RightIdentity (≫) 0 ;
74+
shiftr_0 :: RightIdentity (≫) 0 ;
7575
shiftr_S : ∀ x n, x ≫ n = 2 * x ≫ (1 + n) ∨ x ≫ n = 2 * x ≫ (1 + n) + 1
7676
}.
7777

interfaces/canonical_names.v

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -405,12 +405,12 @@ Notation "f ⁻¹" := (inverse f) (at level 30) : mc_scope.
405405
Class Idempotent `{ea : Equiv A} (f: A → A → A) (x : A) : Prop := idempotency: f x x = x.
406406
Arguments idempotency {A ea} _ _ {Idempotent}.
407407

408-
Class UnaryIdempotent `{Equiv A} (f: A → A) : Prop := unary_idempotent :> Idempotent (@compose A A A) f.
408+
Class UnaryIdempotent `{Equiv A} (f: A → A) : Prop := unary_idempotent :: Idempotent (@compose A A A) f.
409409

410410
Lemma unary_idempotency `{Equiv A} `{!Reflexive (=)} `{!UnaryIdempotent f} x : f (f x) = f x.
411411
Proof. firstorder. Qed.
412412

413-
Class BinaryIdempotent `{Equiv A} (op: A → A → A) : Prop := binary_idempotent :> ∀ x, Idempotent op x.
413+
Class BinaryIdempotent `{Equiv A} (op: A → A → A) : Prop := binary_idempotent :: ∀ x, Idempotent op x.
414414

415415
Class LeftIdentity {A} `{Equiv B} (op : A → B → B) (x : A): Prop := left_identity: ∀ y, op x y = y.
416416
Class RightIdentity `{Equiv A} {B} (op : A → B → A) (y : B): Prop := right_identity: ∀ x, op x y = x.
@@ -428,7 +428,7 @@ Class Commutative `{Equiv B} `(f : A → A → B) : Prop := commutativity: ∀ x
428428
Class HeteroAssociative {A B C AB BC} `{Equiv ABC}
429429
(fA_BC: A → BC → ABC) (fBC: B → C → BC) (fAB_C: AB → C → ABC) (fAB : A → B → AB): Prop
430430
:= associativity : ∀ x y z, fA_BC x (fBC y z) = fAB_C (fAB x y) z.
431-
Class Associative `{Equiv A} f := simple_associativity:> HeteroAssociative f f f f.
431+
Class Associative `{Equiv A} f := simple_associativity:: HeteroAssociative f f f f.
432432
Notation ArrowsAssociative C := (∀ {w x y z: C}, HeteroAssociative (◎) (comp z _ _ ) (◎) (comp y x w)).
433433

434434
Class Involutive `{Equiv A} (f : A → A) := involutive: ∀ x, f (f x) = x.
@@ -451,8 +451,8 @@ Class LeftHeteroDistribute {A B} `{Equiv C} (f : A → B → C) (g_r : B → B
451451
:= distribute_l : ∀ a b c, f a (g_r b c) = g (f a b) (f a c).
452452
Class RightHeteroDistribute {A B} `{Equiv C} (f : A → B → C) (g_l : A → A → A) (g : C → C → C) : Prop
453453
:= distribute_r: ∀ a b c, f (g_l a b) c = g (f a c) (f b c).
454-
Class LeftDistribute`{Equiv A} (f g: A → A → A) := simple_distribute_l :> LeftHeteroDistribute f g g.
455-
Class RightDistribute `{Equiv A} (f g: A → A → A) := simple_distribute_r :> RightHeteroDistribute f g g.
454+
Class LeftDistribute`{Equiv A} (f g: A → A → A) := simple_distribute_l :: LeftHeteroDistribute f g g.
455+
Class RightDistribute `{Equiv A} (f g: A → A → A) := simple_distribute_r :: RightHeteroDistribute f g g.
456456

457457
Class HeteroSymmetric {A} {T : A → A → Type} (R : ∀ {x y}, T x y → T y x → Prop) : Prop :=
458458
hetero_symmetric `(a : T x y) (b : T y x) : R a b → R b a.

interfaces/finite_sets.v

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -42,9 +42,9 @@ Class FSetExtend A `{t : SetType A} :=
4242
Class FSet A `{At : SetType A} `{Ae : Equiv A} `{Ate : SetEquiv A}
4343
`{Aempty : EmptySet A} `{Ajoin : SetJoin A} `{Asingle : SetSingleton A}
4444
`{∀ a₁ a₂ : A, Decision (a₁ = a₂)} `{U : !FSetExtend A} :=
45-
{ fset_bounded_sl :> BoundedJoinSemiLattice (set_type A)
46-
; singleton_mor :> Setoid_Morphism singleton
47-
; fset_extend_mor `{BoundedJoinSemiLattice B} `{!Setoid_Morphism (f : A → B)} :>
45+
{ fset_bounded_sl :: BoundedJoinSemiLattice (set_type A)
46+
; singleton_mor :: Setoid_Morphism singleton
47+
; fset_extend_mor `{BoundedJoinSemiLattice B} `{!Setoid_Morphism (f : A → B)} ::
4848
BoundedJoinSemiLattice_Morphism (fset_extend f)
4949
; fset_extend_correct `{BoundedJoinSemiLattice B} (f : A → B) `{!Setoid_Morphism f} :
5050
f = fset_extend f ∘ singleton
@@ -63,15 +63,15 @@ freely use orders.lattices.alt_Build_JoinSemiLatticeOrder.
6363

6464
Class FSetContainsSpec A `{At : SetType A} `{Ae : Equiv A} `{Ate : SetEquiv A}
6565
`{SetLe A} `{SetContains A} `{Ajoin : SetJoin A} `{Asingle : SetSingleton A} :=
66-
{ fset_join_sl_order :> JoinSemiLatticeOrder (≤)
66+
{ fset_join_sl_order :: JoinSemiLatticeOrder (≤)
6767
; fset_in_singleton_le : ∀ x X, x ∈ X ↔ {{ x }} ≤ X }.
6868

6969
(*
7070
Unfortunately, properties as meet and the differences cannot be uniquely
7171
defined in an algebraic way, therefore we just use set inclusion.
7272
*)
7373
Class FullFSet A {car Ae conAe conAle Acontains Aempty Ajoin Asingle U Adec} `{Adiff : SetDifference A} `{Ameet : SetMeet A} :=
74-
{ full_fset_fset :> @FSet A car Ae conAe Aempty Ajoin Asingle U Adec
75-
; full_fset_contains :> @FSetContainsSpec A car Ae conAe conAle Acontains Ajoin Asingle
74+
{ full_fset_fset :: @FSet A car Ae conAe Aempty Ajoin Asingle U Adec
75+
; full_fset_contains :: @FSetContainsSpec A car Ae conAe conAle Acontains Ajoin Asingle
7676
; fset_in_meet : ∀ X Y x, x ∈ X ⊓ Y ↔ (x ∈ X ∧ x ∈ Y)
7777
; fset_in_difference : ∀ X Y x, x ∈ X∖ Y ↔ (x ∈ X ∧ x ∉ Y) }.

interfaces/functors.v

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@ Section functor_class.
99
Class Functor `(Fmap): Prop :=
1010
{ functor_from: Category C
1111
; functor_to: Category D
12-
; functor_morphism:> ∀ a b: C, Setoid_Morphism (@fmap _ a b)
12+
; functor_morphism:: ∀ a b: C, Setoid_Morphism (@fmap _ a b)
1313
; preserves_id: `(fmap (cat_id: a ⟶ a) = cat_id)
1414
; preserves_comp `(f: y ⟶ z) `(g: x ⟶ y): fmap (f ◎ g) = fmap f ◎ fmap g }.
1515
End functor_class.
@@ -123,8 +123,8 @@ Class SFmap (M : Type → Type) := sfmap: ∀ `(A → B), (M A → M B).
123123

124124
Class SFunctor (M : TypeType)
125125
`{∀ `{Equiv A}, Equiv (M A)} `{SFmap M} : Prop :=
126-
{ sfunctor_setoid `{Setoid A} :> Setoid (M A)
127-
; sfmap_proper `{Setoid A} `{Setoid B} :>
126+
{ sfunctor_setoid `{Setoid A} :: Setoid (M A)
127+
; sfmap_proper `{Setoid A} `{Setoid B} ::
128128
Proper (((=) ==> (=)) ==> ((=) ==> (=))) (@sfmap M _ A B)
129129
; sfmap_id `{Setoid A} : sfmap id = id
130130
; sfmap_comp `{Equiv A} `{Equiv B} `{Equiv C} `{!Setoid_Morphism (f : B → C)} `{!Setoid_Morphism (g : A → B)} :

0 commit comments

Comments
 (0)