Documentation

Orbifolds.Diffeology.Induced

Induced/coinduced diffeologies and inductions/subductions #

In large parts based on Mathlib.Topology.Order.

def DiffeologicalSpace.induced {X : Type u_1} {Y : Type u_2} (f : XY) (dY : DiffeologicalSpace Y) :

The coarsest diffeology on X such that f : X → Y is smooth, also known as the pullback diffeology.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    def DiffeologicalSpace.coinduced {X : Type u_1} {Y : Type u_2} (f : XY) (dX : DiffeologicalSpace X) :

    The finest diffeology on Y such that f : X → Y is smooth, also known as the pushforward diffeology. While this is just the infimum of all diffeologies that make f smooth, we use a slightly more involved construction to make the D-topology defeq to the coinduced topology.

    Equations
    Instances For
      theorem isPlot_induced_iff {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} {n : } {p : Eucl nX} :
      theorem DSmooth.le_induced {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} (hf : DSmooth f) :
      theorem dsmooth_iff_le_induced {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} :
      theorem DiffeologicalSpace.coinduced_eq_sInf {X : Type u_1} {Y : Type u_2} {f : XY} {dX : DiffeologicalSpace X} :

      The coinduced diffeology is the infimum of all diffeologies that make f smooth.

      theorem DSmooth.coinduced_le {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} (hf : DSmooth f) :
      theorem dsmooth_iff_coinduced_le {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} :
      theorem DiffeologicalSpace.coinduced_eq_generateFrom {X : Type u_4} {Y : Type u_5} [dX : DiffeologicalSpace X] {f : XY} :
      coinduced f dX = generateFrom {p : (n : ) × (Eucl nY) | ∃ (q : Eucl p.fstX), IsPlot q p.snd = f q}

      The diffeology coinduced by f : X → Y is generated by all plots of the form f ∘ p for plots p in X.

      theorem dTop_coinduced_comm {X : Type u_4} {Y : Type u_5} {dX : DiffeologicalSpace X} {f : XY} :

      The D-topology of the coinduced diffeology agrees with the coinduced topology.

      theorem DiffeologicalSpace.coinduced_le_iff_le_induced {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} :
      coinduced f dX dY dX induced f dY
      theorem DiffeologicalSpace.induced_mono {X : Type u_1} {Y : Type u_2} {f : XY} {d₁ d₂ : DiffeologicalSpace Y} (h : d₁ d₂) :
      induced f d₁ induced f d₂
      theorem DiffeologicalSpace.coinduced_mono {X : Type u_1} {Y : Type u_2} {f : XY} {d₁ d₂ : DiffeologicalSpace X} (h : d₁ d₂) :
      coinduced f d₁ coinduced f d₂
      @[simp]
      theorem DiffeologicalSpace.induced_top {X : Type u_1} {Y : Type u_2} {f : XY} :
      @[simp]
      theorem DiffeologicalSpace.induced_inf {X : Type u_1} {Y : Type u_2} {d₁ d₂ : DiffeologicalSpace Y} {f : XY} :
      induced f (d₁d₂) = induced f d₁induced f d₂
      @[simp]
      theorem DiffeologicalSpace.induced_iInf {X : Type u_1} {Y : Type u_2} {ι : Sort u_4} {d : ιDiffeologicalSpace Y} {f : XY} :
      induced f (⨅ (i : ι), d i) = ⨅ (i : ι), induced f (d i)
      @[simp]
      theorem DiffeologicalSpace.coinduced_bot {X : Type u_1} {Y : Type u_2} {f : XY} :
      @[simp]
      theorem DiffeologicalSpace.coinduced_sup {X : Type u_1} {Y : Type u_2} {d₁ d₂ : DiffeologicalSpace X} {f : XY} :
      coinduced f (d₁d₂) = coinduced f d₁coinduced f d₂
      @[simp]
      theorem DiffeologicalSpace.coinduced_iSup {X : Type u_1} {Y : Type u_2} {ι : Sort u_4} {d : ιDiffeologicalSpace X} {f : XY} :
      coinduced f (⨆ (i : ι), d i) = ⨆ (i : ι), coinduced f (d i)
      theorem DiffeologicalSpace.induced_compose {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {dZ : DiffeologicalSpace Z} {f : XY} {g : YZ} :
      induced f (induced g dZ) = induced (g f) dZ
      theorem DiffeologicalSpace.induced_const {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {y : Y} :
      induced (fun (x : X) => y) dY =
      theorem DiffeologicalSpace.coinduced_compose {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {dX : DiffeologicalSpace X} {f : XY} {g : YZ} :
      coinduced g (coinduced f dX) = coinduced (g f) dX
      theorem DiffeologicalSpace.coinduced_const {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {y : Y} :
      coinduced (fun (x : X) => y) dX =
      theorem dsmooth_generateFrom_iff {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {g : Set ((n : ) × (Eucl nX))} {f : XY} :
      DSmooth f ∀ (n : ) (p : Eucl nX), n, p gIsPlot (f p)
      theorem dsmooth_induced_dom {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} :
      theorem dsmooth_induced_rng {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {dY : DiffeologicalSpace Y} {dZ : DiffeologicalSpace Z} {f : XY} {g : ZX} :
      theorem dsmooth_coinduced_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} :
      theorem dsmooth_coinduced_dom {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {dX : DiffeologicalSpace X} {dZ : DiffeologicalSpace Z} {f : XY} {g : YZ} :
      theorem dsmooth_le_dom {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {dX' : DiffeologicalSpace X} {f : XY} (h : dX' dX) :
      theorem dsmooth_le_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY dY' : DiffeologicalSpace Y} {f : XY} (h : dY dY') :
      theorem dsmooth_sup_dom {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {dX' : DiffeologicalSpace X} {f : XY} :
      theorem dsmooth_sup_rng_left {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY dY' : DiffeologicalSpace Y} {f : XY} :
      theorem dsmooth_sup_rng_right {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY dY' : DiffeologicalSpace Y} {f : XY} :
      theorem dsmooth_sSup_dom {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} {D : Set (DiffeologicalSpace X)} :
      DSmooth f dD, DSmooth f
      theorem dsmooth_sSup_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} {D : Set (DiffeologicalSpace Y)} (h : dY D) (hf : DSmooth f) :
      theorem dsmooth_iSup_dom {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} {ι : Type u_4} {D : ιDiffeologicalSpace X} :
      DSmooth f ∀ (i : ι), DSmooth f
      theorem dsmooth_iSup_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} {ι : Type u_4} {D : ιDiffeologicalSpace Y} {i : ι} (h : DSmooth f) :
      theorem dsmooth_inf_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY dY' : DiffeologicalSpace Y} {f : XY} :
      theorem dsmooth_inf_dom_left {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {dX' : DiffeologicalSpace X} {f : XY} :
      theorem dsmooth_inf_dom_right {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {dX' : DiffeologicalSpace X} {f : XY} :
      theorem dsmooth_sInf_dom {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {dY : DiffeologicalSpace Y} {f : XY} {D : Set (DiffeologicalSpace X)} (h : dX D) :
      theorem dsmooth_sInf_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} {D : Set (DiffeologicalSpace Y)} :
      DSmooth f dD, DSmooth f
      theorem dsmooth_iInf_dom {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} {ι : Type u_4} {D : ιDiffeologicalSpace X} {i : ι} :
      theorem dsmooth_iInf_rng {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} {ι : Type u_4} {D : ιDiffeologicalSpace Y} :
      DSmooth f ∀ (i : ι), DSmooth f
      theorem dsmooth_bot {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} :
      theorem dsmooth_top {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} :
      theorem dsmooth_id_iff_le {X : Type u_1} {dX dX' : DiffeologicalSpace X} :
      DSmooth id dX dX'
      theorem dsmooth_id_of_le {X : Type u_1} {dX dX' : DiffeologicalSpace X} (h : dX dX') :
      theorem dTop_induced_le_induced_dTop {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} :

      The D-topology of the induced diffeology is always at least as fine as the induced topology.

      theorem dTop_induced_comm {X : Type u_4} {Y : Type u_5} {dY : DiffeologicalSpace Y} {f : XY} (hf : IsOpen (Set.range f)) :

      For functions whose range is D-open, the D-topology of the induced diffeology agrees with the induced topology.

      structure IsDInducing {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) :

      We call a map "D-inducing" (short for diffeologically inducing) if its domain carries the diffeology induced by it. This doesn't seem to appear in the literature, but we introduce it as a weaker version of inductions for situations where injectivity isn't actually needed.

      This is analogous to inducing maps in topology, whereas inductions are analagous to embeddings.

      Instances For
        theorem isDInducing_iff {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) :
        structure IsInduction {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) extends IsDInducing f :

        An induction is a map between diffeological spaces that is both inducing and injective. This is analogous to embeddings in topology.

        Instances For
          theorem isInduction_iff {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) :
          structure IsDCoinducing {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) :

          We call a map "D-coinducing" (short for diffeologically coinducing) if its codomain carries the diffeology coinduced by it. This doesn't seem to appear in the literature, but we introduce it as a weaker version of subductions for situations where surjectivity isn't actually needed.

          This is analogous to coinducing maps in topology, whereas subductions are analagous to quotient maps.

          Instances For
            theorem isDCoinducing_iff {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) :
            structure IsSubduction {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) extends IsDCoinducing f :

            A subduction is a map between diffeological spaces that is both coinducing and surjective. This is analogous to quotient maps in topology.

            Instances For
              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      theorem isDInducing_induced {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} :
                      theorem isDCoinducing_coinduced {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} :
                      theorem Function.Injective.isInduction_induced {X : Type u_1} {Y : Type u_2} {dY : DiffeologicalSpace Y} {f : XY} (hf : Injective f) :
                      theorem Function.Surjective.isSubduction_coinduced {X : Type u_1} {Y : Type u_2} {dX : DiffeologicalSpace X} {f : XY} (hf : Surjective f) :
                      theorem IsDInducing.dsmooth {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} (hf : IsDInducing f) :
                      theorem IsDCoinducing.dsmooth {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} (hf : IsDCoinducing f) :
                      theorem IsDInducing.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDInducing g) (hf : IsDInducing f) :
                      theorem IsInduction.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsInduction g) (hf : IsInduction f) :
                      theorem IsDInducing.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDInducing g) (hf : IsDInducing f) :
                      IsDInducing fun (x : X) => g (f x)
                      theorem IsInduction.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsInduction g) (hf : IsInduction f) :
                      IsInduction fun (x : X) => g (f x)
                      theorem IsDInducing.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDInducing g) (h : IsDInducing (g f)) :
                      theorem IsInduction.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsInduction g) (h : IsInduction (g f)) :
                      theorem IsDInducing.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDInducing g) :
                      theorem IsInduction.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsInduction g) :
                      theorem IsDInducing.of_comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : DSmooth f) (hg : DSmooth g) (h : IsDInducing (g f)) :
                      theorem IsInduction.of_comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : DSmooth f) (hg : DSmooth g) (h : IsInduction (g f)) :
                      theorem IsDCoinducing.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDCoinducing g) (hf : IsDCoinducing f) :
                      theorem IsSubduction.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsSubduction g) (hf : IsSubduction f) :
                      theorem IsDCoinducing.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDCoinducing g) (hf : IsDCoinducing f) :
                      IsDCoinducing fun (x : X) => g (f x)
                      theorem IsSubduction.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsSubduction g) (hf : IsSubduction f) :
                      IsSubduction fun (x : X) => g (f x)
                      theorem IsDCoinducing.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : IsDCoinducing f) (h : IsDCoinducing (g f)) :
                      theorem IsSubduction.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : IsSubduction f) (h : IsSubduction (g f)) :
                      theorem IsDCoinducing.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : IsDCoinducing f) :
                      theorem IsSubduction.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : IsSubduction f) :
                      theorem IsDCoinducing.of_comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : DSmooth f) (hg : DSmooth g) (h : IsDCoinducing (g f)) :
                      theorem IsSubduction.of_comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : DSmooth f) (hg : DSmooth g) (h : IsSubduction (g f)) :
                      theorem IsDInducing.dsmooth_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsDInducing g) :
                      theorem IsDInducing.isPlot_iff {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} {n : } {p : Eucl nX} (hf : IsDInducing f) :
                      theorem IsDCoinducing.dsmooth_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hf : IsDCoinducing f) :
                      theorem Function.LeftInverse.isInduction {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} {g : YX} (h : LeftInverse f g) (hf : DSmooth f) (hg : DSmooth g) :
                      theorem Function.LeftInverse.isSubduction {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} {g : YX} (h : LeftInverse f g) (hf : DSmooth f) (hg : DSmooth g) :
                      structure IsOpenInduction {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) extends IsInduction f :

                      An open induction is an induction that is also an open map with respect to the D-topologies. What makes this an especially interesting notion compared to e.g. closed inductions is that while inductions are not always embeddings, open inductions are always open embeddings; furthermore, inductions are open iff their range is open. In comparison, an induction whose range is closed can still be both not an embedding and not a closed map.

                      Instances For
                        theorem isOpenInduction_iff {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] (f : XY) :
                        theorem IsOpenInduction.isEmbedding {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} (hf : IsOpenInduction f) :
                        theorem IsOpenInduction.isOpen_range {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} (hf : IsOpenInduction f) :
                        theorem IsInduction.isOpenInduction_of_isOpen_range {X : Type u_1} {Y : Type u_2} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] {f : XY} (hf : IsInduction f) (hf' : IsOpen (Set.range f)) :
                        theorem IsOpenInduction.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsOpenInduction g) (hf : IsOpenInduction f) :
                        theorem IsOpenInduction.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsOpenInduction g) (hf : IsOpenInduction f) :
                        IsOpenInduction fun (x : X) => g (f x)
                        theorem IsOpenInduction.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsOpenInduction g) (h : IsOpenInduction (g f)) :
                        theorem IsOpenInduction.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsOpenInduction g) :
                        theorem IsOpenInduction.of_comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [dX : DiffeologicalSpace X] [dY : DiffeologicalSpace Y] [dZ : DiffeologicalSpace Z] {f : XY} {g : YZ} (hg : IsInduction g) (h : IsOpenInduction (g f)) :