@@ -383,3 +383,43 @@ module _ {L₁ : RawLoop a ℓ₁}
383383 { isLoopMonomorphism = isLoopMonomorphism F.isLoopMonomorphism G.isLoopMonomorphism
384384 ; surjective = Func.surjective _ _ (_≈_ L₃) F.surjective G.surjective
385385 } where module F = IsLoopIsomorphism f-iso; module G = IsLoopIsomorphism g-iso
386+
387+ ------------------------------------------------------------------------
388+ -- KleeneAlgebra
389+
390+ module _ {K₁ : RawKleeneAlgebra a ℓ₁}
391+ {K₂ : RawKleeneAlgebra b ℓ₂}
392+ {K₃ : RawKleeneAlgebra c ℓ₃}
393+ (open RawKleeneAlgebra )
394+ (≈₃-trans : Transitive (_≈_ K₃))
395+ {f : Carrier K₁ → Carrier K₂}
396+ {g : Carrier K₂ → Carrier K₃}
397+ where
398+
399+ isKleeneAlgebraHomomorphism
400+ : IsKleeneAlgebraHomomorphism K₁ K₂ f
401+ → IsKleeneAlgebraHomomorphism K₂ K₃ g
402+ → IsKleeneAlgebraHomomorphism K₁ K₃ (g ∘ f)
403+ isKleeneAlgebraHomomorphism f-homo g-homo = record
404+ { isSemiringHomomorphism = isSemiringHomomorphism ≈₃-trans F.isSemiringHomomorphism G.isSemiringHomomorphism
405+ ; ⋆-homo = λ x → ≈₃-trans (G.⟦⟧-cong (F.⋆-homo x)) (G.⋆-homo (f x))
406+ } where module F = IsKleeneAlgebraHomomorphism f-homo; module G = IsKleeneAlgebraHomomorphism g-homo
407+
408+ isKleeneAlgebraMonomorphism
409+ : IsKleeneAlgebraMonomorphism K₁ K₂ f
410+ → IsKleeneAlgebraMonomorphism K₂ K₃ g
411+ → IsKleeneAlgebraMonomorphism K₁ K₃ (g ∘ f)
412+ isKleeneAlgebraMonomorphism f-mono g-mono = record
413+ { isKleeneAlgebraHomomorphism = isKleeneAlgebraHomomorphism F.isKleeneAlgebraHomomorphism G.isKleeneAlgebraHomomorphism
414+ ; injective = F.injective ∘ G.injective
415+ } where module F = IsKleeneAlgebraMonomorphism f-mono; module G = IsKleeneAlgebraMonomorphism g-mono
416+
417+ isKleeneAlgebraIsomorphism
418+ : IsKleeneAlgebraIsomorphism K₁ K₂ f
419+ → IsKleeneAlgebraIsomorphism K₂ K₃ g
420+ → IsKleeneAlgebraIsomorphism K₁ K₃ (g ∘ f)
421+ isKleeneAlgebraIsomorphism f-iso g-iso = record
422+ { isKleeneAlgebraMonomorphism = isKleeneAlgebraMonomorphism F.isKleeneAlgebraMonomorphism G.isKleeneAlgebraMonomorphism
423+ ; surjective = Func.surjective (_≈_ K₁) (_≈_ K₂) (_≈_ K₃) F.surjective G.surjective
424+ } where module F = IsKleeneAlgebraIsomorphism f-iso; module G = IsKleeneAlgebraIsomorphism g-iso
425+
0 commit comments