src/HOL/Analysis/Analysis.thy
author nipkow
Fri, 28 Dec 2018 10:29:59 +0100
changeset 69517 dc20f278e8f3
parent 69144 f13b82281715
child 69676 56acd449da41
permissions -rw-r--r--
tuned style and headers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     1
theory Analysis
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents: 56370
diff changeset
     2
imports
63626
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
     3
  Lebesgue_Integral_Substitution
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66277
diff changeset
     4
  Improper_Integral
63626
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
     5
  Embed_Measure
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
     6
  Complete_Measure
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
     7
  Radon_Nikodym
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
     8
  Fashoda_Theorem
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents: 56370
diff changeset
     9
  Determinants
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents: 68000
diff changeset
    10
  Cross3
63129
paulson <lp15@cam.ac.uk>
parents: 63078
diff changeset
    11
  Homeomorphism
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents: 56370
diff changeset
    12
  Bounded_Continuous_Function
69144
f13b82281715 new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents: 68625
diff changeset
    13
  Abstract_Topology
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents: 64006
diff changeset
    14
  Function_Topology
66277
512b0dc09061 HOL-Analysis: Infinite products
eberlm <eberlm@in.tum.de>
parents: 65040
diff changeset
    15
  Infinite_Products
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66296
diff changeset
    16
  Infinite_Set_Sum
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
    17
  Weierstrass_Theorems
63078
e49dc94eb624 Theory of polyhedra: faces, extreme points, polytopes, and the Krein–Milman
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
    18
  Polytope
64846
de4e3df6693d Jordan Curve Theorem
paulson <lp15@cam.ac.uk>
parents: 64790
diff changeset
    19
  Jordan_Curve
65039
87972e6177bc New theory about Winding Numbers
paulson <lp15@cam.ac.uk>
parents: 64846
diff changeset
    20
  Winding_Numbers
66826
0d60d2118544 Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
paulson <lp15@cam.ac.uk>
parents: 66480
diff changeset
    21
  Riemann_Mapping
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
    22
  Poly_Roots
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62049
diff changeset
    23
  Conformal_Mappings
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66296
diff changeset
    24
  FPS_Convergence
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61518
diff changeset
    25
  Generalised_Binomial_Theorem
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
    26
  Gamma_Function
68000
40b790c5a11d Oops! Change_Of_Vars was not being imported to Analysis!
paulson <lp15@cam.ac.uk>
parents: 67996
diff changeset
    27
  Change_Of_Vars
67727
ce3e87a51488 moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents: 67278
diff changeset
    28
  Lipschitz
68625
2ec84498f562 HOL-Analysis: Volume of a simplex, Heron's theorem
Manuel Eberl <eberlm@in.tum.de>
parents: 68465
diff changeset
    29
  Simplex_Content
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    30
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    31
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    32
end