src/HOL/Analysis/Analysis.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 79857 819c28a7280f
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
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
69676
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
     2
  imports
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
     3
  (* Linear Algebra *)
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
     4
  Convex
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
     5
  Determinants
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
     6
  (* Topology *)
77939
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents: 77102
diff changeset
     7
  FSigma
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents: 77102
diff changeset
     8
  Sum_Topology
77941
c35f06b0931e new material ported from HOL Light's metric.ml
paulson <lp15@cam.ac.uk>
parents: 77939
diff changeset
     9
  Abstract_Topological_Spaces
78093
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77941
diff changeset
    10
  Abstract_Metric_Spaces
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
    11
  Urysohn
69676
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    12
  Connected
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69676
diff changeset
    13
  Abstract_Limits
77102
780161d4b55c Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents: 74475
diff changeset
    14
  Isolated
79857
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
    15
  Sparse_In
69676
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    16
  (* Functional Analysis *)
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    17
  Elementary_Normed_Spaces
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    18
  Norm_Arith
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    19
  (* Vector Analysis *)
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    20
  Convex_Euclidean_Space
71194
26b35a97bddb tuned manual
nipkow
parents: 71191
diff changeset
    21
  Operator_Norm
26b35a97bddb tuned manual
nipkow
parents: 71191
diff changeset
    22
  (* Unsorted *)
26b35a97bddb tuned manual
nipkow
parents: 71191
diff changeset
    23
  Line_Segment
26b35a97bddb tuned manual
nipkow
parents: 71191
diff changeset
    24
  Derivative
26b35a97bddb tuned manual
nipkow
parents: 71191
diff changeset
    25
  Cartesian_Euclidean_Space
78890
d8045bc0544e Added Kronecker's approximation theorem. Requires adding Real_Asymp to HOL-Analysis. Funny syntax issue in Probability/Projective_Family
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
    26
  Kronecker_Approximation_Theorem
71194
26b35a97bddb tuned manual
nipkow
parents: 71191
diff changeset
    27
  Weierstrass_Theorems
69676
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    28
  (* Measure and Integration Theory *)
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    29
  Ball_Volume
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    30
  Integral_Test
56acd449da41 chapters for analysis manual
immler
parents: 69144
diff changeset
    31
  Improper_Integral
70694
ae37b8fbf023 New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents: 70621
diff changeset
    32
  Equivalence_Measurable_On_Borel
63626
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
    33
  Lebesgue_Integral_Substitution
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
    34
  Embed_Measure
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
    35
  Complete_Measure
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63594
diff changeset
    36
  Radon_Nikodym
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
    37
  Fashoda_Theorem
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents: 68000
diff changeset
    38
  Cross3
63129
paulson <lp15@cam.ac.uk>
parents: 63078
diff changeset
    39
  Homeomorphism
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents: 56370
diff changeset
    40
  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
    41
  Abstract_Topology
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
    42
  Product_Topology
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
    43
  Lindelof_Spaces
66277
512b0dc09061 HOL-Analysis: Infinite products
eberlm <eberlm@in.tum.de>
parents: 65040
diff changeset
    44
  Infinite_Products
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents: 71200
diff changeset
    45
  Infinite_Sum
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66296
diff changeset
    46
  Infinite_Set_Sum
63078
e49dc94eb624 Theory of polyhedra: faces, extreme points, polytopes, and the Krein–Milman
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
    47
  Polytope
64846
de4e3df6693d Jordan Curve Theorem
paulson <lp15@cam.ac.uk>
parents: 64790
diff changeset
    48
  Jordan_Curve
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
    49
  Poly_Roots
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61518
diff changeset
    50
  Generalised_Binomial_Theorem
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63129
diff changeset
    51
  Gamma_Function
68000
40b790c5a11d Oops! Change_Of_Vars was not being imported to Analysis!
paulson <lp15@cam.ac.uk>
parents: 67996
diff changeset
    52
  Change_Of_Vars
70621
1afcfb7fdff4 entry point for analysis without integration theory
immler
parents: 70178
diff changeset
    53
  Multivariate_Analysis
68625
2ec84498f562 HOL-Analysis: Volume of a simplex, Heron's theorem
Manuel Eberl <eberlm@in.tum.de>
parents: 68465
diff changeset
    54
  Simplex_Content
71189
954ee5acaae0 Split off new HOL-Complex_Analysis session from HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 70694
diff changeset
    55
  FPS_Convergence
954ee5acaae0 Split off new HOL-Complex_Analysis session from HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 70694
diff changeset
    56
  Smooth_Paths
71197
36961c681fed tuned Analysis manual
immler
parents: 71194
diff changeset
    57
  Abstract_Euclidean_Space
71200
3548d54ce3ee split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents: 71197
diff changeset
    58
  Function_Metric
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    59
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    60
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    61
end