Documentation

CombinatorialGames.Mathlib.Concept

Formal concept analysis #

This file is a near-copy of Mathlib.Order.Concept, but uses the updated names and has some extra lemmas.

Once Mathlib and this repository update, remove this file!

Lower and upper polars #

def upperPolar {α : Type u_2} {β : Type u_3} (r : αβProp) (s : Set α) :
Set β

The upper polar of s : Set α along a relation r : α → β → Prop is the set of all elements which r relates to all elements of s.

Equations
Instances For
    def lowerPolar {α : Type u_2} {β : Type u_3} (r : αβProp) (t : Set β) :
    Set α

    The lower polar of t : Set β along a relation r : α → β → Prop is the set of all elements which r relates to all elements of t.

    Equations
    Instances For
      theorem subset_upperPolar_iff_subset_lowerPolar {α : Type u_2} {β : Type u_3} {r : αβProp} {s : Set α} {t : Set β} :
      theorem upperPolar_swap {α : Type u_2} {β : Type u_3} (r : αβProp) (t : Set β) :
      theorem lowerPolar_swap {α : Type u_2} {β : Type u_3} (r : αβProp) (s : Set α) :
      @[simp]
      theorem upperPolar_empty {α : Type u_2} {β : Type u_3} (r : αβProp) :
      @[simp]
      theorem lowerPolar_empty {α : Type u_2} {β : Type u_3} (r : αβProp) :
      @[simp]
      theorem upperPolar_union {α : Type u_2} {β : Type u_3} (r : αβProp) (s₁ s₂ : Set α) :
      upperPolar r (s₁ s₂) = upperPolar r s₁ upperPolar r s₂
      @[simp]
      theorem lowerPolar_union {α : Type u_2} {β : Type u_3} (r : αβProp) (t₁ t₂ : Set β) :
      lowerPolar r (t₁ t₂) = lowerPolar r t₁ lowerPolar r t₂
      @[simp]
      theorem upperPolar_iUnion {ι : Sort u_1} {α : Type u_2} {β : Type u_3} (r : αβProp) (f : ιSet α) :
      upperPolar r (⋃ (i : ι), f i) = ⋂ (i : ι), upperPolar r (f i)
      @[simp]
      theorem lowerPolar_iUnion {ι : Sort u_1} {α : Type u_2} {β : Type u_3} (r : αβProp) (f : ιSet β) :
      lowerPolar r (⋃ (i : ι), f i) = ⋂ (i : ι), lowerPolar r (f i)
      theorem upperPolar_iUnion₂ {ι : Sort u_1} {α : Type u_2} {β : Type u_3} {κ : ιSort u_4} (r : αβProp) (f : (i : ι) → κ iSet α) :
      upperPolar r (⋃ (i : ι), ⋃ (j : κ i), f i j) = ⋂ (i : ι), ⋂ (j : κ i), upperPolar r (f i j)
      theorem lowerPolar_iUnion₂ {ι : Sort u_1} {α : Type u_2} {β : Type u_3} {κ : ιSort u_4} (r : αβProp) (f : (i : ι) → κ iSet β) :
      lowerPolar r (⋃ (i : ι), ⋃ (j : κ i), f i j) = ⋂ (i : ι), ⋂ (j : κ i), lowerPolar r (f i j)
      theorem subset_lowerPolar_upperPolar {α : Type u_2} {β : Type u_3} (r : αβProp) (s : Set α) :
      theorem subset_upperPolar_lowerPolar {α : Type u_2} {β : Type u_3} (r : αβProp) (t : Set β) :
      @[simp]
      theorem upperPolar_lowerPolar_upperPolar {α : Type u_2} {β : Type u_3} (r : αβProp) (s : Set α) :
      @[simp]
      theorem lowerPolar_upperPolar_lowerPolar {α : Type u_2} {β : Type u_3} (r : αβProp) (t : Set β) :
      theorem upperPolar_anti {α : Type u_2} {β : Type u_3} (r : αβProp) :
      theorem lowerPolar_anti {α : Type u_2} {β : Type u_3} (r : αβProp) :

      Concepts #

      structure Concept (α : Type u_2) (β : Type u_3) (r : αβProp) :
      Type (max u_2 u_3)

      The formal concepts of a relation. A concept of r : α → β → Prop is a pair of sets s, t such that s is the set of all elements that are r-related to all of t and t is the set of all elements that are r-related to all of s.

      • extent : Set α

        The extent of a concept.

      • intent : Set β

        The intent of a concept.

      • upperPolar_extent : upperPolar r self.extent = self.intent

        The intent consists of all elements related to all elements of the extent.

      • lowerPolar_intent : lowerPolar r self.intent = self.extent

        The extent consists of all elements related to all elements of the intent.

      Instances For
        theorem Concept.ext {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} (h : c.extent = d.extent) :
        c = d
        theorem Concept.ext_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
        c = d c.extent = d.extent
        theorem Concept.ext' {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} (h : c.intent = d.intent) :
        c = d
        theorem Concept.extent_injective {α : Type u_2} {β : Type u_3} {r : αβProp} :
        theorem Concept.intent_injective {α : Type u_2} {β : Type u_3} {r : αβProp} :
        def Concept.copy {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) (e : Set α) (he : e = c.extent) (i : Set β) (hi : i = c.intent) :
        Concept α β r

        Copy a concept, adjusting definitional equalities.

        Equations
        • c.copy e he i hi = { extent := e, intent := i, upperPolar_extent := , lowerPolar_intent := }
        Instances For
          @[simp]
          theorem Concept.extent_copy {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) (e : Set α) (he : e = c.extent) (i : Set β) (hi : i = c.intent) :
          (c.copy e he i hi).extent = e
          @[simp]
          theorem Concept.intent_copy {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) (e : Set α) (he : e = c.extent) (i : Set β) (hi : i = c.intent) :
          (c.copy e he i hi).intent = i
          theorem Concept.copy_eq {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) (e : Set α) (he : e = c.extent) (i : Set β) (hi : i = c.intent) :
          c.copy e he i hi = c
          theorem Concept.rel_extent_intent {α : Type u_2} {β : Type u_3} {r : αβProp} {c : Concept α β r} {x : α} {y : β} (hx : x c.extent) (hy : y c.intent) :
          r x y
          theorem Concept.disjoint_extent_intent {α : Type u_2} {r' : ααProp} [IsIrrefl α r'] (c' : Concept α α r') :

          Note that if r' is the relation, this theorem will often not be true!

          theorem Concept.mem_extent_of_rel_extent {α : Type u_2} {r' : ααProp} {c' : Concept α α r'} [IsTrans α r'] {x y : α} (hy : r' y x) (hx : x c'.extent) :
          y c'.extent
          theorem Concept.mem_intent_of_intent_rel {α : Type u_2} {r' : ααProp} {c' : Concept α α r'} [IsTrans α r'] {x y : α} (hy : r' x y) (hx : x c'.intent) :
          y c'.intent
          theorem Concept.codisjoint_extent_intent {α : Type u_2} {r' : ααProp} [IsTrichotomous α r'] [IsTrans α r'] (c' : Concept α α r') :
          theorem Concept.isCompl_extent_intent {α : Type u_2} {r' : ααProp} [IsStrictTotalOrder α r'] (c' : Concept α α r') :
          theorem Concept.isLowerSet_extent {α : Type u_5} [Preorder α] (c : Concept α α fun (x1 x2 : α) => x1 x2) :

          See isLowerSet_extent' for the theorem on a <-concept.

          theorem Concept.isUpperSet_intent {α : Type u_5} [Preorder α] (c : Concept α α fun (x1 x2 : α) => x1 x2) :

          See isUpperSet_intent' for the theorem on a <-concept.

          theorem Concept.isLowerSet_extent' {α : Type u_5} [PartialOrder α] (c : Concept α α fun (x1 x2 : α) => x1 < x2) :

          See isLowerSet_extent for the theorem on a -concept.

          theorem Concept.isUpperSet_intent' {α : Type u_5} [PartialOrder α] (c : Concept α α fun (x1 x2 : α) => x1 < x2) :

          See isUpperSet_intent for the theorem on a -concept.

          instance Concept.instSupConcept {α : Type u_2} {β : Type u_3} {r : αβProp} :
          Max (Concept α β r)
          Equations
          • One or more equations did not get rendered due to their size.
          instance Concept.instInfConcept {α : Type u_2} {β : Type u_3} {r : αβProp} :
          Min (Concept α β r)
          Equations
          • One or more equations did not get rendered due to their size.
          @[simp]
          theorem Concept.extent_subset_extent_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
          @[simp]
          theorem Concept.extent_ssubset_extent_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
          @[simp]
          theorem Concept.intent_subset_intent_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
          @[simp]
          theorem Concept.intent_ssubset_intent_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
          theorem Concept.strictMono_extent {α : Type u_2} {β : Type u_3} {r : αβProp} :
          theorem Concept.strictAnti_intent {α : Type u_2} {β : Type u_3} {r : αβProp} :
          instance Concept.instLatticeConcept {α : Type u_2} {β : Type u_3} {r : αβProp} :
          Lattice (Concept α β r)
          Equations
          • One or more equations did not get rendered due to their size.
          instance Concept.instBoundedOrderConcept {α : Type u_2} {β : Type u_3} {r : αβProp} :
          Equations
          • One or more equations did not get rendered due to their size.
          instance Concept.instSupSet {α : Type u_2} {β : Type u_3} {r : αβProp} :
          SupSet (Concept α β r)
          Equations
          • One or more equations did not get rendered due to their size.
          instance Concept.instInfSet {α : Type u_2} {β : Type u_3} {r : αβProp} :
          InfSet (Concept α β r)
          Equations
          • One or more equations did not get rendered due to their size.
          instance Concept.instCompleteLattice {α : Type u_2} {β : Type u_3} {r : αβProp} :
          Equations
          • One or more equations did not get rendered due to their size.
          @[simp]
          theorem Concept.extent_top {α : Type u_2} {β : Type u_3} {r : αβProp} :
          @[simp]
          theorem Concept.intent_top {α : Type u_2} {β : Type u_3} {r : αβProp} :
          @[simp]
          theorem Concept.extent_bot {α : Type u_2} {β : Type u_3} {r : αβProp} :
          @[simp]
          theorem Concept.intent_bot {α : Type u_2} {β : Type u_3} {r : αβProp} :
          @[simp]
          theorem Concept.extent_sup {α : Type u_2} {β : Type u_3} {r : αβProp} (c d : Concept α β r) :
          (cd).extent = lowerPolar r (c.intent d.intent)
          @[simp]
          theorem Concept.intent_sup {α : Type u_2} {β : Type u_3} {r : αβProp} (c d : Concept α β r) :
          (cd).intent = c.intent d.intent
          @[simp]
          theorem Concept.extent_inf {α : Type u_2} {β : Type u_3} {r : αβProp} (c d : Concept α β r) :
          (cd).extent = c.extent d.extent
          @[simp]
          theorem Concept.intent_inf {α : Type u_2} {β : Type u_3} {r : αβProp} (c d : Concept α β r) :
          (cd).intent = upperPolar r (c.extent d.extent)
          @[simp]
          theorem Concept.extent_sSup {α : Type u_2} {β : Type u_3} {r : αβProp} (S : Set (Concept α β r)) :
          (sSup S).extent = lowerPolar r (⋂ cS, c.intent)
          @[simp]
          theorem Concept.intent_sSup {α : Type u_2} {β : Type u_3} {r : αβProp} (S : Set (Concept α β r)) :
          (sSup S).intent = cS, c.intent
          @[simp]
          theorem Concept.extent_sInf {α : Type u_2} {β : Type u_3} {r : αβProp} (S : Set (Concept α β r)) :
          (sInf S).extent = cS, c.extent
          @[simp]
          theorem Concept.intent_sInf {α : Type u_2} {β : Type u_3} {r : αβProp} (S : Set (Concept α β r)) :
          (sInf S).intent = upperPolar r (⋂ cS, c.extent)
          instance Concept.instInhabited {α : Type u_2} {β : Type u_3} {r : αβProp} :
          Inhabited (Concept α β r)
          Equations
          def Concept.swap {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) :

          Swap the sets of a concept to make it a concept of the dual context.

          Equations
          • c.swap = { extent := c.intent, intent := c.extent, upperPolar_extent := , lowerPolar_intent := }
          Instances For
            @[simp]
            theorem Concept.swap_intent {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) :
            @[simp]
            theorem Concept.swap_extent {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) :
            @[simp]
            theorem Concept.swap_swap {α : Type u_2} {β : Type u_3} {r : αβProp} (c : Concept α β r) :
            c.swap.swap = c
            @[simp]
            theorem Concept.swap_le_swap_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
            c.swap d.swap d c
            @[simp]
            theorem Concept.swap_lt_swap_iff {α : Type u_2} {β : Type u_3} {r : αβProp} {c d : Concept α β r} :
            c.swap < d.swap d < c
            def Concept.swapEquiv {α : Type u_2} {β : Type u_3} {r : αβProp} :

            The dual of a concept lattice is isomorphic to the concept lattice of the dual context.

            Equations
            Instances For
              @[simp]
              theorem Concept.swapEquiv_symm_apply {α : Type u_2} {β : Type u_3} {r : αβProp} (a✝ : Concept β α (Function.swap r)) :
              @[simp]
              theorem Concept.swapEquiv_apply {α : Type u_2} {β : Type u_3} {r : αβProp} (a✝ : (Concept α β r)ᵒᵈ) :