File: ua_subalgebra.v

package info (click to toggle)
coq-math-classes 9.0.0-2
  • links: PTS, VCS
  • area: main
  • in suites: experimental
  • size: 1,120 kB
  • sloc: python: 22; makefile: 21; sh: 2
file content (101 lines) | stat: -rw-r--r-- 3,203 bytes parent folder | download | duplicates (3)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
Require Import
  Coq.Classes.RelationClasses
  MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.interfaces.canonical_names MathClasses.theory.categories MathClasses.interfaces.abstract_algebra.
Require
  MathClasses.categories.algebras MathClasses.theory.forget_algebra.

Section subalgebras.
  Context `{Algebra sign A} (P: ∀ s, A s → Prop).

  (* We begin by describing what it means for P to be a proper closed subset: *)

  Fixpoint op_closed {o: OpType (sorts sign)}: op_type A o → Prop :=
    match o with
    | ne_list.one x => P x
    | ne_list.cons _ _ => λ d, ∀ z, P _ z → op_closed (d z)
    end.

  Global Instance op_closed_proper: ∀ `{∀ s, Proper ((=) ==> iff) (P s)} o, Proper ((=) ==> iff) (@op_closed o).
  Proof with intuition.
   induction o; simpl; intros x y E.
    rewrite E...
   split; intros.
    apply (IHo (x z))... apply E...
   apply (IHo _ (y z))... apply E...
  Qed.

  Class ClosedSubset: Prop :=
    { subset_proper:: ∀ s, Proper ((=) ==> iff) (P s)
    ; subset_closed: ∀ o, op_closed (algebra_op o) }.

  (* Now suppose P is closed in this way. *)

  Context `{ClosedSubset}.

  (* Our new algebra's elements are just those for which P holds: *)

  Definition carrier s := sig (P s).

  Hint Extern 4 => progress unfold carrier: typeclass_instances.

  (* We can implement closed operations in the new algebra: *)

  Program Fixpoint close_op {d}: ∀ (o: op_type A d), op_closed o → op_type carrier d :=
    match d with
    | ne_list.one _ => λ o c, exist _ o (c)
    | ne_list.cons _ _ => λ o c X, close_op (o X) (c X (proj2_sig X))
    end.

  Global Instance impl: AlgebraOps sign carrier := λ o, close_op (algebra_op o) (subset_closed o).

  (* By showing that these ops are proper, we get our new algebra: *)

  Definition close_op_proper d (o0 o1: op_type A d)
    (P': op_closed o0) (Q: op_closed o1): o0 = o1 → close_op o0 P' = close_op o1 Q.
  Proof with intuition.
   induction d; simpl in *...
   intros [x p] [y q] U.
   apply (IHd _ _ (P' x p) (Q y q)), H2...
  Qed.

  Global Instance subalgebra: Algebra sign carrier.
  Proof. constructor. apply _. intro. apply close_op_proper, algebra_propers. Qed.

  (* And we have the obvious projection morphism: *)

  Definition proj s := @proj1_sig (A s) (P s).

  Global Instance: HomoMorphism sign carrier A proj.
  Proof with try apply _.
   constructor...
    constructor...
    firstorder.
   intro.
   unfold impl, algebra_op.
   generalize (subset_closed o).
   unfold algebra_op.
   generalize (H o).
   induction (sign o); simpl; intuition.
  Qed.

  (* Which is mono because proj is injective. *)
  Instance: `{Injective (proj i)}.
  Proof.
    split.
     firstorder.
    apply (@homo_proper sign carrier A
      (fun s : sorts sign => @sig_equiv (A s) (e s) (P s)) _ _ _ _).
    apply _.
  Qed.

  Global Instance: Mono (algebras.arrow _ proj).
  Proof.
   apply forget_algebra.mono.
   apply categories.product.mono.
   intro. apply setoids.mono.
   simpl. apply _.
  Qed. (* this really should be completely automatic. *)
End subalgebras.

#[global]
Hint Extern 10 (Equiv (carrier _ _)) => apply @sig_equiv : typeclass_instances.