Documentation

Mathlib.Combinatorics.Enumerative.Composition

Compositions #

A composition of a natural number n is a decomposition n = i₀ + ... + i_{k-1} of n into a sum of positive integers. Combinatorially, it corresponds to a decomposition of {0, ..., n-1} into non-empty blocks of consecutive integers, where the iⱼ are the lengths of the blocks. This notion is closely related to that of a partition of n, but in a composition of n the order of the iⱼs matters.

We implement two different structures covering these two viewpoints on compositions. The first one, made of a list of positive integers summing to n, is the main one and is called Composition n. The second one is useful for combinatorial arguments (for instance to show that the number of compositions of n is 2^(n-1)). It is given by a subset of {0, ..., n} containing 0 and n, where the elements of the subset (other than n) correspond to the leftmost points of each block. The main API is built on Composition n, and we provide an equivalence between the two types.

Main functions #

Let c : Composition n be a composition of n. Then

Compositions can also be used to split lists. Let l be a list of length n and c a composition of n.

We turn to the second viewpoint on compositions, that we realize as a finset of Fin (n+1). c : CompositionAsSet n is a structure made of a finset of Fin (n+1) called c.boundaries and proofs that it contains 0 and n. (Taking a finset of Fin n containing 0 would not make sense in the edge case n = 0, while the previous description works in all cases). The elements of this set (other than n) correspond to leftmost points of blocks. Thus, there is an equiv between Composition n and CompositionAsSet n. We only construct basic API on CompositionAsSet (notably c.length and c.blocks) to be able to construct this equiv, called compositionEquiv n. Since there is a straightforward equiv between CompositionAsSet n and finsets of {1, ..., n-1} (obtained by removing 0 and n from a CompositionAsSet and called compositionAsSetEquiv n), we deduce that CompositionAsSet n and Composition n are both fintypes of cardinality 2^(n - 1) (see compositionAsSet_card and composition_card).

Implementation details #

The main motivation for this structure and its API is in the construction of the composition of formal multilinear series, and the proof that the composition of analytic functions is analytic.

The representation of a composition as a list is very handy as lists are very flexible and already have a well-developed API.

Tags #

Composition, partition

References #

theorem Composition.ext {n : } (x : Composition n) (y : Composition n) (blocks : x.blocks = y.blocks) :
x = y
theorem Composition.ext_iff {n : } (x : Composition n) (y : Composition n) :
x = y x.blocks = y.blocks
structure Composition (n : ) :

A composition of n is a list of positive integers summing to n.

  • blocks : List

    List of positive integers summing to n

  • blocks_pos : ∀ {i : }, i self.blocks0 < i

    Proof of positivity for blocks

  • blocks_sum : self.blocks.sum = n

    Proof that blocks sums to n

Instances For
    theorem Composition.blocks_pos {n : } (self : Composition n) {i : } :
    i self.blocks0 < i

    Proof of positivity for blocks

    theorem Composition.blocks_sum {n : } (self : Composition n) :
    self.blocks.sum = n

    Proof that blocks sums to n

    theorem CompositionAsSet.ext_iff {n : } (x : CompositionAsSet n) (y : CompositionAsSet n) :
    x = y x.boundaries = y.boundaries
    theorem CompositionAsSet.ext {n : } (x : CompositionAsSet n) (y : CompositionAsSet n) (boundaries : x.boundaries = y.boundaries) :
    x = y
    structure CompositionAsSet (n : ) :

    Combinatorial viewpoint on a composition of n, by seeing it as non-empty blocks of consecutive integers in {0, ..., n-1}. We register every block by its left end-point, yielding a finset containing 0. As this does not make sense for n = 0, we add n to this finset, and get a finset of {0, ..., n} containing 0 and n. This is the data in the structure CompositionAsSet n.

    • boundaries : Finset (Fin n.succ)

      Combinatorial viewpoint on a composition of n as consecutive integers {0, ..., n-1}

    • zero_mem : 0 self.boundaries

      Proof that 0 is a member of boundaries

    • getLast_mem : Fin.last n self.boundaries

      Last element of the composition

    Instances For
      theorem CompositionAsSet.zero_mem {n : } (self : CompositionAsSet n) :
      0 self.boundaries

      Proof that 0 is a member of boundaries

      theorem CompositionAsSet.getLast_mem {n : } (self : CompositionAsSet n) :
      Fin.last n self.boundaries

      Last element of the composition

      Equations
      • instInhabitedCompositionAsSet = { default := { boundaries := Finset.univ, zero_mem := , getLast_mem := } }

      Compositions #

      A composition of an integer n is a decomposition n = i₀ + ... + i_{k-1} of n into a sum of positive integers.

      Equations
      @[reducible]

      The length of a composition, i.e., the number of blocks in the composition.

      Equations
      • c.length = c.blocks.length
      Instances For
        theorem Composition.blocks_length {n : } (c : Composition n) :
        c.blocks.length = c.length
        def Composition.blocksFun {n : } (c : Composition n) :
        Fin c.length

        The blocks of a composition, seen as a function on Fin c.length. When composing analytic functions using compositions, this is the main player.

        Equations
        • c.blocksFun = c.blocks.get
        Instances For
          theorem Composition.ofFn_blocksFun {n : } (c : Composition n) :
          List.ofFn c.blocksFun = c.blocks
          theorem Composition.sum_blocksFun {n : } (c : Composition n) :
          (Finset.univ.sum fun (i : Fin c.length) => c.blocksFun i) = n
          theorem Composition.blocksFun_mem_blocks {n : } (c : Composition n) (i : Fin c.length) :
          c.blocksFun i c.blocks
          @[simp]
          theorem Composition.one_le_blocks {n : } (c : Composition n) {i : } (h : i c.blocks) :
          1 i
          @[simp]
          theorem Composition.one_le_blocks' {n : } (c : Composition n) {i : } (h : i < c.length) :
          1 c.blocks.get i, h
          @[simp]
          theorem Composition.blocks_pos' {n : } (c : Composition n) (i : ) (h : i < c.length) :
          0 < c.blocks.get i, h
          theorem Composition.one_le_blocksFun {n : } (c : Composition n) (i : Fin c.length) :
          1 c.blocksFun i
          theorem Composition.length_le {n : } (c : Composition n) :
          c.length n
          theorem Composition.length_pos_of_pos {n : } (c : Composition n) (h : 0 < n) :
          0 < c.length
          def Composition.sizeUpTo {n : } (c : Composition n) (i : ) :

          The sum of the sizes of the blocks in a composition up to i.

          Equations
          Instances For
            @[simp]
            theorem Composition.sizeUpTo_zero {n : } (c : Composition n) :
            c.sizeUpTo 0 = 0
            theorem Composition.sizeUpTo_ofLength_le {n : } (c : Composition n) (i : ) (h : c.length i) :
            c.sizeUpTo i = n
            @[simp]
            theorem Composition.sizeUpTo_length {n : } (c : Composition n) :
            c.sizeUpTo c.length = n
            theorem Composition.sizeUpTo_le {n : } (c : Composition n) (i : ) :
            c.sizeUpTo i n
            theorem Composition.sizeUpTo_succ {n : } (c : Composition n) {i : } (h : i < c.length) :
            c.sizeUpTo (i + 1) = c.sizeUpTo i + c.blocks.get i, h
            theorem Composition.sizeUpTo_succ' {n : } (c : Composition n) (i : Fin c.length) :
            c.sizeUpTo (i + 1) = c.sizeUpTo i + c.blocksFun i
            theorem Composition.sizeUpTo_strict_mono {n : } (c : Composition n) {i : } (h : i < c.length) :
            c.sizeUpTo i < c.sizeUpTo (i + 1)
            def Composition.boundary {n : } (c : Composition n) :
            Fin (c.length + 1) ↪o Fin (n + 1)

            The i-th boundary of a composition, i.e., the leftmost point of the i-th block. We include a virtual point at the right of the last block, to make for a nice equiv with CompositionAsSet n.

            Equations
            Instances For
              @[simp]
              theorem Composition.boundary_zero {n : } (c : Composition n) :
              c.boundary 0 = 0
              @[simp]
              theorem Composition.boundary_last {n : } (c : Composition n) :
              c.boundary (Fin.last c.length) = Fin.last n
              def Composition.boundaries {n : } (c : Composition n) :
              Finset (Fin (n + 1))

              The boundaries of a composition, i.e., the leftmost point of all the blocks. We include a virtual point at the right of the last block, to make for a nice equiv with CompositionAsSet n.

              Equations
              • c.boundaries = Finset.map c.boundary.toEmbedding Finset.univ
              Instances For
                theorem Composition.card_boundaries_eq_succ_length {n : } (c : Composition n) :
                c.boundaries.card = c.length + 1

                To c : Composition n, one can associate a CompositionAsSet n by registering the leftmost point of each block, and adding a virtual point at the right of the last block.

                Equations
                • c.toCompositionAsSet = { boundaries := c.boundaries, zero_mem := , getLast_mem := }
                Instances For
                  theorem Composition.orderEmbOfFin_boundaries {n : } (c : Composition n) :
                  c.boundaries.orderEmbOfFin = c.boundary

                  The canonical increasing bijection between Fin (c.length + 1) and c.boundaries is exactly c.boundary.

                  def Composition.embedding {n : } (c : Composition n) (i : Fin c.length) :
                  Fin (c.blocksFun i) ↪o Fin n

                  Embedding the i-th block of a composition (identified with Fin (c.blocks_fun i)) into Fin n at the relevant position.

                  Equations
                  Instances For
                    @[simp]
                    theorem Composition.coe_embedding {n : } (c : Composition n) (i : Fin c.length) (j : Fin (c.blocksFun i)) :
                    ((c.embedding i) j) = c.sizeUpTo i + j
                    theorem Composition.index_exists {n : } (c : Composition n) {j : } (h : j < n) :
                    ∃ (i : ), j < c.sizeUpTo (i + 1) i < c.length

                    index_exists asserts there is some i with j < c.size_up_to (i+1). In the next definition index we use Nat.find to produce the minimal such index.

                    def Composition.index {n : } (c : Composition n) (j : Fin n) :
                    Fin c.length

                    c.index j is the index of the block in the composition c containing j.

                    Equations
                    Instances For
                      theorem Composition.lt_sizeUpTo_index_succ {n : } (c : Composition n) (j : Fin n) :
                      j < c.sizeUpTo (c.index j).succ
                      theorem Composition.sizeUpTo_index_le {n : } (c : Composition n) (j : Fin n) :
                      c.sizeUpTo (c.index j) j
                      def Composition.invEmbedding {n : } (c : Composition n) (j : Fin n) :
                      Fin (c.blocksFun (c.index j))

                      Mapping an element j of Fin n to the element in the block containing it, identified with Fin (c.blocks_fun (c.index j)) through the canonical increasing bijection.

                      Equations
                      • c.invEmbedding j = j - c.sizeUpTo (c.index j),
                      Instances For
                        @[simp]
                        theorem Composition.coe_invEmbedding {n : } (c : Composition n) (j : Fin n) :
                        (c.invEmbedding j) = j - c.sizeUpTo (c.index j)
                        theorem Composition.embedding_comp_inv {n : } (c : Composition n) (j : Fin n) :
                        (c.embedding (c.index j)) (c.invEmbedding j) = j
                        theorem Composition.mem_range_embedding_iff {n : } (c : Composition n) {j : Fin n} {i : Fin c.length} :
                        j Set.range (c.embedding i) c.sizeUpTo i j j < c.sizeUpTo (i).succ
                        theorem Composition.disjoint_range {n : } (c : Composition n) {i₁ : Fin c.length} {i₂ : Fin c.length} (h : i₁ i₂) :
                        Disjoint (Set.range (c.embedding i₁)) (Set.range (c.embedding i₂))

                        The embeddings of different blocks of a composition are disjoint.

                        theorem Composition.mem_range_embedding {n : } (c : Composition n) (j : Fin n) :
                        j Set.range (c.embedding (c.index j))
                        theorem Composition.mem_range_embedding_iff' {n : } (c : Composition n) {j : Fin n} {i : Fin c.length} :
                        j Set.range (c.embedding i) i = c.index j
                        theorem Composition.index_embedding {n : } (c : Composition n) (i : Fin c.length) (j : Fin (c.blocksFun i)) :
                        c.index ((c.embedding i) j) = i
                        theorem Composition.invEmbedding_comp {n : } (c : Composition n) (i : Fin c.length) (j : Fin (c.blocksFun i)) :
                        (c.invEmbedding ((c.embedding i) j)) = j
                        def Composition.blocksFinEquiv {n : } (c : Composition n) :
                        (i : Fin c.length) × Fin (c.blocksFun i) Fin n

                        Equivalence between the disjoint union of the blocks (each of them seen as Fin (c.blocks_fun i)) with Fin n.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          theorem Composition.blocksFun_congr {n₁ : } {n₂ : } (c₁ : Composition n₁) (c₂ : Composition n₂) (i₁ : Fin c₁.length) (i₂ : Fin c₂.length) (hn : n₁ = n₂) (hc : c₁.blocks = c₂.blocks) (hi : i₁ = i₂) :
                          c₁.blocksFun i₁ = c₂.blocksFun i₂
                          theorem Composition.sigma_eq_iff_blocks_eq {c : (n : ) × Composition n} {c' : (n : ) × Composition n} :
                          c = c' c.snd.blocks = c'.snd.blocks

                          Two compositions (possibly of different integers) coincide if and only if they have the same sequence of blocks.

                          The composition Composition.ones #

                          The composition made of blocks all of size 1.

                          Equations
                          Instances For
                            Equations
                            @[simp]
                            theorem Composition.ones_length (n : ) :
                            (Composition.ones n).length = n
                            @[simp]
                            @[simp]
                            theorem Composition.ones_blocksFun (n : ) (i : Fin (Composition.ones n).length) :
                            (Composition.ones n).blocksFun i = 1
                            @[simp]
                            theorem Composition.ones_sizeUpTo (n : ) (i : ) :
                            (Composition.ones n).sizeUpTo i = min i n
                            @[simp]
                            theorem Composition.ones_embedding {n : } (i : Fin (Composition.ones n).length) (h : 0 < (Composition.ones n).blocksFun i) :
                            ((Composition.ones n).embedding i) 0, h = i,
                            theorem Composition.eq_ones_iff {n : } {c : Composition n} :
                            c = Composition.ones n ic.blocks, i = 1
                            theorem Composition.ne_ones_iff {n : } {c : Composition n} :
                            c Composition.ones n ic.blocks, 1 < i

                            The composition Composition.single #

                            def Composition.single (n : ) (h : 0 < n) :

                            The composition made of a single block of size n.

                            Equations
                            Instances For
                              @[simp]
                              theorem Composition.single_length {n : } (h : 0 < n) :
                              (Composition.single n h).length = 1
                              @[simp]
                              theorem Composition.single_blocks {n : } (h : 0 < n) :
                              (Composition.single n h).blocks = [n]
                              @[simp]
                              theorem Composition.single_blocksFun {n : } (h : 0 < n) (i : Fin (Composition.single n h).length) :
                              (Composition.single n h).blocksFun i = n
                              @[simp]
                              theorem Composition.single_embedding {n : } (h : 0 < n) (i : Fin n) :
                              ((Composition.single n h).embedding 0) i = i
                              theorem Composition.eq_single_iff_length {n : } (h : 0 < n) {c : Composition n} :
                              c = Composition.single n h c.length = 1
                              theorem Composition.ne_single_iff {n : } (hn : 0 < n) {c : Composition n} :
                              c Composition.single n hn ∀ (i : Fin c.length), c.blocksFun i < n

                              Splitting a list #

                              Given a list of length n and a composition c of n, one can split l into c.length sublists of respective lengths c.blocks_fun 0, ..., c.blocks_fun (c.length-1). This is inverse to the join operation.

                              def List.splitWrtCompositionAux {α : Type u_1} :
                              List αList List (List α)

                              Auxiliary for List.splitWrtComposition.

                              Equations
                              • x.splitWrtCompositionAux [] = []
                              • x.splitWrtCompositionAux (n :: ns) = match List.splitAt n x with | (l₁, l₂) => l₁ :: l₂.splitWrtCompositionAux ns
                              Instances For
                                def List.splitWrtComposition {n : } {α : Type u_1} (l : List α) (c : Composition n) :
                                List (List α)

                                Given a list of length n and a composition [i₁, ..., iₖ] of n, split l into a list of k lists corresponding to the blocks of the composition, of respective lengths i₁, ..., iₖ. This makes sense mostly when n = l.length, but this is not necessary for the definition.

                                Equations
                                • l.splitWrtComposition c = l.splitWrtCompositionAux c.blocks
                                Instances For
                                  theorem List.splitWrtCompositionAux_cons {α : Type u_1} (l : List α) (n : ) (ns : List ) :
                                  l.splitWrtCompositionAux (n :: ns) = List.take n l :: (List.drop n l).splitWrtCompositionAux ns
                                  theorem List.length_splitWrtCompositionAux {α : Type u_1} (l : List α) (ns : List ) :
                                  (l.splitWrtCompositionAux ns).length = ns.length
                                  @[simp]
                                  theorem List.length_splitWrtComposition {n : } {α : Type u_1} (l : List α) (c : Composition n) :
                                  (l.splitWrtComposition c).length = c.length

                                  When one splits a list along a composition c, the number of sublists thus created is c.length.

                                  theorem List.map_length_splitWrtCompositionAux {α : Type u_1} {ns : List } {l : List α} :
                                  ns.sum l.lengthList.map List.length (l.splitWrtCompositionAux ns) = ns
                                  theorem List.map_length_splitWrtComposition {α : Type u_1} (l : List α) (c : Composition l.length) :
                                  List.map List.length (l.splitWrtComposition c) = c.blocks

                                  When one splits a list along a composition c, the lengths of the sublists thus created are given by the block sizes in c.

                                  theorem List.length_pos_of_mem_splitWrtComposition {α : Type u_1} {l : List α} {l' : List α} {c : Composition l.length} (h : l' l.splitWrtComposition c) :
                                  0 < l'.length
                                  theorem List.sum_take_map_length_splitWrtComposition {α : Type u_1} (l : List α) (c : Composition l.length) (i : ) :
                                  (List.take i (List.map List.length (l.splitWrtComposition c))).sum = c.sizeUpTo i
                                  theorem List.get_splitWrtCompositionAux {α : Type u_1} (l : List α) (ns : List ) {i : } (hi : i < (l.splitWrtCompositionAux ns).length) :
                                  (l.splitWrtCompositionAux ns).get i, hi = List.drop (List.take i ns).sum (List.take (List.take (i + 1) ns).sum l)
                                  theorem List.get_splitWrtComposition' {n : } {α : Type u_1} (l : List α) (c : Composition n) {i : } (hi : i < (l.splitWrtComposition c).length) :
                                  (l.splitWrtComposition c).get i, hi = List.drop (c.sizeUpTo i) (List.take (c.sizeUpTo (i + 1)) l)

                                  The i-th sublist in the splitting of a list l along a composition c, is the slice of l between the indices c.sizeUpTo i and c.sizeUpTo (i+1), i.e., the indices in the i-th block of the composition.

                                  theorem List.get_splitWrtComposition {n : } {α : Type u_1} (l : List α) (c : Composition n) (i : Fin (l.splitWrtComposition c).length) :
                                  (l.splitWrtComposition c).get i = List.drop (c.sizeUpTo i) (List.take (c.sizeUpTo (i + 1)) l)
                                  theorem List.join_splitWrtCompositionAux {α : Type u_1} {ns : List } {l : List α} :
                                  ns.sum = l.length(l.splitWrtCompositionAux ns).join = l
                                  @[simp]
                                  theorem List.join_splitWrtComposition {α : Type u_1} (l : List α) (c : Composition l.length) :
                                  (l.splitWrtComposition c).join = l

                                  If one splits a list along a composition, and then joins the sublists, one gets back the original list.

                                  @[simp]
                                  theorem List.splitWrtComposition_join {α : Type u_1} (L : List (List α)) (c : Composition L.join.length) (h : List.map List.length L = c.blocks) :
                                  L.join.splitWrtComposition c = L

                                  If one joins a list of lists and then splits the join along the right composition, one gets back the original list of lists.

                                  Compositions as sets #

                                  Combinatorial viewpoints on compositions, seen as finite subsets of Fin (n+1) containing 0 and n, where the points of the set (other than n) correspond to the leftmost points of each block.

                                  Bijection between compositions of n and subsets of {0, ..., n-2}, defined by considering the restriction of the subset to {1, ..., n-1} and shifting to the left by one.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    theorem CompositionAsSet.boundaries_nonempty {n : } (c : CompositionAsSet n) :
                                    c.boundaries.Nonempty
                                    theorem CompositionAsSet.card_boundaries_pos {n : } (c : CompositionAsSet n) :
                                    0 < c.boundaries.card

                                    Number of blocks in a CompositionAsSet.

                                    Equations
                                    • c.length = c.boundaries.card - 1
                                    Instances For
                                      theorem CompositionAsSet.card_boundaries_eq_succ_length {n : } (c : CompositionAsSet n) :
                                      c.boundaries.card = c.length + 1
                                      theorem CompositionAsSet.length_lt_card_boundaries {n : } (c : CompositionAsSet n) :
                                      c.length < c.boundaries.card
                                      theorem CompositionAsSet.lt_length {n : } (c : CompositionAsSet n) (i : Fin c.length) :
                                      i + 1 < c.boundaries.card
                                      theorem CompositionAsSet.lt_length' {n : } (c : CompositionAsSet n) (i : Fin c.length) :
                                      i < c.boundaries.card
                                      def CompositionAsSet.boundary {n : } (c : CompositionAsSet n) :
                                      Fin c.boundaries.card ↪o Fin (n + 1)

                                      Canonical increasing bijection from Fin c.boundaries.card to c.boundaries.

                                      Equations
                                      • c.boundary = c.boundaries.orderEmbOfFin
                                      Instances For
                                        @[simp]
                                        theorem CompositionAsSet.boundary_zero {n : } (c : CompositionAsSet n) :
                                        c.boundary 0, = 0
                                        @[simp]
                                        theorem CompositionAsSet.boundary_length {n : } (c : CompositionAsSet n) :
                                        c.boundary c.length, = Fin.last n
                                        def CompositionAsSet.blocksFun {n : } (c : CompositionAsSet n) (i : Fin c.length) :

                                        Size of the i-th block in a CompositionAsSet, seen as a function on Fin c.length.

                                        Equations
                                        • c.blocksFun i = (c.boundary i + 1, ) - (c.boundary i, )
                                        Instances For
                                          theorem CompositionAsSet.blocksFun_pos {n : } (c : CompositionAsSet n) (i : Fin c.length) :
                                          0 < c.blocksFun i

                                          List of the sizes of the blocks in a CompositionAsSet.

                                          Equations
                                          Instances For
                                            @[simp]
                                            theorem CompositionAsSet.blocks_length {n : } (c : CompositionAsSet n) :
                                            c.blocks.length = c.length
                                            theorem CompositionAsSet.blocks_partial_sum {n : } (c : CompositionAsSet n) {i : } (h : i < c.boundaries.card) :
                                            (List.take i c.blocks).sum = (c.boundary i, h)
                                            theorem CompositionAsSet.mem_boundaries_iff_exists_blocks_sum_take_eq {n : } (c : CompositionAsSet n) {j : Fin (n + 1)} :
                                            j c.boundaries i < c.boundaries.card, (List.take i c.blocks).sum = j
                                            theorem CompositionAsSet.blocks_sum {n : } (c : CompositionAsSet n) :
                                            c.blocks.sum = n

                                            Associating a Composition n to a CompositionAsSet n, by registering the sizes of the blocks as a list of positive integers.

                                            Equations
                                            • c.toComposition = { blocks := c.blocks, blocks_pos := , blocks_sum := }
                                            Instances For

                                              Equivalence between compositions and compositions as sets #

                                              In this section, we explain how to go back and forth between a Composition and a CompositionAsSet, by showing that their blocks and length and boundaries correspond to each other, and construct an equivalence between them called compositionEquiv.

                                              @[simp]
                                              theorem Composition.toCompositionAsSet_length {n : } (c : Composition n) :
                                              c.toCompositionAsSet.length = c.length
                                              @[simp]
                                              theorem CompositionAsSet.toComposition_length {n : } (c : CompositionAsSet n) :
                                              c.toComposition.length = c.length
                                              @[simp]
                                              theorem Composition.toCompositionAsSet_blocks {n : } (c : Composition n) :
                                              c.toCompositionAsSet.blocks = c.blocks
                                              @[simp]
                                              theorem CompositionAsSet.toComposition_blocks {n : } (c : CompositionAsSet n) :
                                              c.toComposition.blocks = c.blocks
                                              @[simp]
                                              theorem CompositionAsSet.toComposition_boundaries {n : } (c : CompositionAsSet n) :
                                              c.toComposition.boundaries = c.boundaries
                                              @[simp]
                                              theorem Composition.toCompositionAsSet_boundaries {n : } (c : Composition n) :
                                              c.toCompositionAsSet.boundaries = c.boundaries

                                              Equivalence between Composition n and CompositionAsSet n.

                                              Equations
                                              Instances For