diff --git a/doc/editors/web.md b/doc/editors/web.md --- a/doc/editors/web.md +++ b/doc/editors/web.md @@ -1,20 +1,25 @@ Web site & Site generator ------------------------- The web site is generated by a set of Python scripts. As input, they take - the metadata of the entries (as specified in the `metadata` folder), - dependencies between entries (as generated by the `afp_dependencies` tool), and - static templates (in `metadata/templates`) The output will be written to the `web` folder, which is supposed to be committed into the repository. It can be inspected by opening any of the contained HTML files in the browser. The script can be invoked without any arguments: ./admin/sitegen Changing static content, e.g. the submission guidelines, works by editing the appropriate template file (see above) and re-running the `sitegen` script. + +To publish website changes without publishing a new entry, you can use the +`publish` script with the `-` argument: + + ./admin/publish - diff --git a/metadata/metadata b/metadata/metadata --- a/metadata/metadata +++ b/metadata/metadata @@ -1,11330 +1,11404 @@ [Arith_Prog_Rel_Primes] title = Arithmetic progressions and relative primes author = José Manuel Rodríguez Caballero topic = Mathematics/Number theory date = 2020-02-01 notify = jose.manuel.rodriguez.caballero@ut.ee abstract = This article provides a formalization of the solution obtained by the author of the Problem “ARITHMETIC PROGRESSIONS” from the Putnam exam problems of 2002. The statement of the problem is as follows: For which integers n > 1 does the set of positive integers less than and relatively prime to n constitute an arithmetic progression? [Banach_Steinhaus] title = Banach-Steinhaus Theorem author = Dominique Unruh , Jose Manuel Rodriguez Caballero topic = Mathematics/Analysis date = 2020-05-02 notify = jose.manuel.rodriguez.caballero@ut.ee, unruh@ut.ee abstract = We formalize in Isabelle/HOL a result due to S. Banach and H. Steinhaus known as the Banach-Steinhaus theorem or Uniform boundedness principle: a pointwise-bounded family of continuous linear operators from a Banach space to a normed space is uniformly bounded. Our approach is an adaptation to Isabelle/HOL of a proof due to A. Sokal. [Complex_Geometry] title = Complex Geometry author = Filip Marić , Danijela Simić topic = Mathematics/Geometry date = 2019-12-16 notify = danijela@matf.bg.ac.rs, filip@matf.bg.ac.rs, boutry@unistra.fr abstract = A formalization of geometry of complex numbers is presented. Fundamental objects that are investigated are the complex plane extended by a single infinite point, its objects (points, lines and circles), and groups of transformations that act on them (e.g., inversions and Möbius transformations). Most objects are defined algebraically, but correspondence with classical geometric definitions is shown. [Poincare_Disc] title = Poincaré Disc Model author = Danijela Simić , Filip Marić , Pierre Boutry topic = Mathematics/Geometry date = 2019-12-16 notify = danijela@matf.bg.ac.rs, filip@matf.bg.ac.rs, boutry@unistra.fr abstract = We describe formalization of the Poincaré disc model of hyperbolic geometry within the Isabelle/HOL proof assistant. The model is defined within the extended complex plane (one dimensional complex projectives space ℂP1), formalized in the AFP entry “Complex Geometry”. Points, lines, congruence of pairs of points, betweenness of triples of points, circles, and isometries are defined within the model. It is shown that the model satisfies all Tarski's axioms except the Euclid's axiom. It is shown that it satisfies its negation and the limiting parallels axiom (which proves it to be a model of hyperbolic geometry). [Fourier] title = Fourier Series author = Lawrence C Paulson topic = Mathematics/Analysis date = 2019-09-06 notify = lp15@cam.ac.uk abstract = This development formalises the square integrable functions over the reals and the basics of Fourier series. It culminates with a proof that every well-behaved periodic function can be approximated by a Fourier series. The material is ported from HOL Light: https://github.com/jrh13/hol-light/blob/master/100/fourier.ml [Generic_Deriving] title = Deriving generic class instances for datatypes author = Jonas Rädle , Lars Hupel topic = Computer science/Data structures date = 2018-11-06 notify = jonas.raedle@gmail.com abstract =

We provide a framework for automatically deriving instances for generic type classes. Our approach is inspired by Haskell's generic-deriving package and Scala's shapeless library. In addition to generating the code for type class functions, we also attempt to automatically prove type class laws for these instances. As of now, however, some manual proofs are still required for recursive datatypes.

Note: There are already articles in the AFP that provide automatic instantiation for a number of classes. Concretely, Deriving allows the automatic instantiation of comparators, linear orders, equality, and hashing. Show instantiates a Haskell-style show class.

Our approach works for arbitrary classes (with some Isabelle/HOL overhead for each class), but a smaller set of datatypes.

[Partial_Order_Reduction] title = Partial Order Reduction author = Julian Brunner topic = Computer science/Automata and formal languages date = 2018-06-05 notify = brunnerj@in.tum.de abstract = This entry provides a formalization of the abstract theory of ample set partial order reduction. The formalization includes transition systems with actions, trace theory, as well as basics on finite, infinite, and lazy sequences. We also provide a basic framework for static analysis on concurrent systems with respect to the ample set condition. [CakeML] title = CakeML author = Lars Hupel , Yu Zhang <> contributors = Johannes Åman Pohjola <> topic = Computer science/Programming languages/Language definitions date = 2018-03-12 notify = hupel@in.tum.de abstract = CakeML is a functional programming language with a proven-correct compiler and runtime system. This entry contains an unofficial version of the CakeML semantics that has been exported from the Lem specifications to Isabelle. Additionally, there are some hand-written theory files that adapt the exported code to Isabelle and port proofs from the HOL4 formalization, e.g. termination and equivalence proofs. [CakeML_Codegen] title = A Verified Code Generator from Isabelle/HOL to CakeML author = Lars Hupel topic = Computer science/Programming languages/Compiling, Logic/Rewriting date = 2019-07-08 notify = lars@hupel.info abstract = This entry contains the formalization that accompanies my PhD thesis (see https://lars.hupel.info/research/codegen/). I develop a verified compilation toolchain from executable specifications in Isabelle/HOL to CakeML abstract syntax trees. This improves over the state-of-the-art in Isabelle by providing a trustworthy procedure for code generation. [DiscretePricing] title = Pricing in discrete financial models author = Mnacho Echenim topic = Mathematics/Probability theory, Mathematics/Games and economics date = 2018-07-16 notify = mnacho.echenim@univ-grenoble-alpes.fr abstract = We have formalized the computation of fair prices for derivative products in discrete financial models. As an application, we derive a way to compute fair prices of derivative products in the Cox-Ross-Rubinstein model of a financial market, thus completing the work that was presented in this paper. extra-history = Change history: [2019-05-12]: Renamed discr_mkt predicate to stk_strict_subs and got rid of predicate A for a more natural definition of the type discrete_market; renamed basic quantity processes for coherent notation; renamed value_process into val_process and closing_value_process to cls_val_process; relaxed hypothesis of lemma CRR_market_fair_price. Added functions to price some basic options. (revision 0b813a1a833f)
[Pell] title = Pell's Equation author = Manuel Eberl topic = Mathematics/Number theory date = 2018-06-23 notify = manuel@pruvisto.org abstract =

This article gives the basic theory of Pell's equation x2 = 1 + Dy2, where D ∈ ℕ is a parameter and x, y are integer variables.

The main result that is proven is the following: If D is not a perfect square, then there exists a fundamental solution (x0, y0) that is not the trivial solution (1, 0) and which generates all other solutions (x, y) in the sense that there exists some n ∈ ℕ such that |x| + |y| √D = (x0 + y0 √D)n. This also implies that the set of solutions is infinite, and it gives us an explicit and executable characterisation of all the solutions.

Based on this, simple executable algorithms for computing the fundamental solution and the infinite sequence of all non-negative solutions are also provided.

[WebAssembly] title = WebAssembly author = Conrad Watt topic = Computer science/Programming languages/Language definitions date = 2018-04-29 notify = caw77@cam.ac.uk abstract = This is a mechanised specification of the WebAssembly language, drawn mainly from the previously published paper formalisation of Haas et al. Also included is a full proof of soundness of the type system, together with a verified type checker and interpreter. We include only a partial procedure for the extraction of the type checker and interpreter here. For more details, please see our paper in CPP 2018. [Knuth_Morris_Pratt] title = The string search algorithm by Knuth, Morris and Pratt author = Fabian Hellauer , Peter Lammich topic = Computer science/Algorithms date = 2017-12-18 notify = hellauer@in.tum.de, lammich@in.tum.de abstract = The Knuth-Morris-Pratt algorithm is often used to show that the problem of finding a string s in a text t can be solved deterministically in O(|s| + |t|) time. We use the Isabelle Refinement Framework to formulate and verify the algorithm. Via refinement, we apply some optimisations and finally use the Sepref tool to obtain executable code in Imperative/HOL. [Minkowskis_Theorem] title = Minkowski's Theorem author = Manuel Eberl topic = Mathematics/Geometry, Mathematics/Number theory date = 2017-07-13 notify = manuel@pruvisto.org abstract =

Minkowski's theorem relates a subset of ℝn, the Lebesgue measure, and the integer lattice ℤn: It states that any convex subset of ℝn with volume greater than 2n contains at least one lattice point from ℤn\{0}, i. e. a non-zero point with integer coefficients.

A related theorem which directly implies this is Blichfeldt's theorem, which states that any subset of ℝn with a volume greater than 1 contains two different points whose difference vector has integer components.

The entry contains a proof of both theorems.

[Name_Carrying_Type_Inference] title = Verified Metatheory and Type Inference for a Name-Carrying Simply-Typed Lambda Calculus author = Michael Rawson topic = Computer science/Programming languages/Type systems date = 2017-07-09 notify = mr644@cam.ac.uk, michaelrawson76@gmail.com abstract = I formalise a Church-style simply-typed \(\lambda\)-calculus, extended with pairs, a unit value, and projection functions, and show some metatheory of the calculus, such as the subject reduction property. Particular attention is paid to the treatment of names in the calculus. A nominal style of binding is used, but I use a manual approach over Nominal Isabelle in order to extract an executable type inference algorithm. More information can be found in my undergraduate dissertation. [Propositional_Proof_Systems] title = Propositional Proof Systems author = Julius Michaelis , Tobias Nipkow topic = Logic/Proof theory date = 2017-06-21 notify = maintainafpppt@liftm.de abstract = We formalize a range of proof systems for classical propositional logic (sequent calculus, natural deduction, Hilbert systems, resolution) and prove the most important meta-theoretic results about semantics and proofs: compactness, soundness, completeness, translations between proof systems, cut-elimination, interpolation and model existence. [Optics] title = Optics author = Simon Foster , Frank Zeyda topic = Computer science/Functional programming, Mathematics/Algebra date = 2017-05-25 notify = simon.foster@york.ac.uk abstract = Lenses provide an abstract interface for manipulating data types through spatially-separated views. They are defined abstractly in terms of two functions, get, the return a value from the source type, and put that updates the value. We mechanise the underlying theory of lenses, in terms of an algebraic hierarchy of lenses, including well-behaved and very well-behaved lenses, each lens class being characterised by a set of lens laws. We also mechanise a lens algebra in Isabelle that enables their composition and comparison, so as to allow construction of complex lenses. This is accompanied by a large library of algebraic laws. Moreover we also show how the lens classes can be applied by instantiating them with a number of Isabelle data types. extra-history = Change history: [2020-03-02]: Added partial bijective and symmetric lenses. Improved alphabet command generating additional lenses and results. Several additional lens relations, including observational equivalence. Additional theorems throughout. Adaptations for Isabelle 2020. (revision 44e2e5c) [2021-01-27] Addition of new theorems throughout, particularly for prisms. New "chantype" command allows the definition of an algebraic datatype with generated prisms. New "dataspace" command allows the definition of a local-based state space, including lenses and prisms. Addition of various examples for the above. (revision 89cf045a) [Game_Based_Crypto] title = Game-based cryptography in HOL author = Andreas Lochbihler , S. Reza Sefidgar <>, Bhargav Bhatt topic = Computer science/Security/Cryptography date = 2017-05-05 notify = mail@andreas-lochbihler.de abstract =

In this AFP entry, we show how to specify game-based cryptographic security notions and formally prove secure several cryptographic constructions from the literature using the CryptHOL framework. Among others, we formalise the notions of a random oracle, a pseudo-random function, an unpredictable function, and of encryption schemes that are indistinguishable under chosen plaintext and/or ciphertext attacks. We prove the random-permutation/random-function switching lemma, security of the Elgamal and hashed Elgamal public-key encryption scheme and correctness and security of several constructions with pseudo-random functions.

Our proofs follow the game-hopping style advocated by Shoup and Bellare and Rogaway, from which most of the examples have been taken. We generalise some of their results such that they can be reused in other proofs. Thanks to CryptHOL's integration with Isabelle's parametricity infrastructure, many simple hops are easily justified using the theory of representation independence.

extra-history = Change history: [2018-09-28]: added the CryptHOL tutorial for game-based cryptography (revision 489a395764ae) [Multi_Party_Computation] title = Multi-Party Computation author = David Aspinall , David Butler topic = Computer science/Security date = 2019-05-09 notify = dbutler@turing.ac.uk abstract = We use CryptHOL to consider Multi-Party Computation (MPC) protocols. MPC was first considered by Yao in 1983 and recent advances in efficiency and an increased demand mean it is now deployed in the real world. Security is considered using the real/ideal world paradigm. We first define security in the semi-honest security setting where parties are assumed not to deviate from the protocol transcript. In this setting we prove multiple Oblivious Transfer (OT) protocols secure and then show security for the gates of the GMW protocol. We then define malicious security, this is a stronger notion of security where parties are assumed to be fully corrupted by an adversary. In this setting we again consider OT, as it is a fundamental building block of almost all MPC protocols. [Sigma_Commit_Crypto] title = Sigma Protocols and Commitment Schemes author = David Butler , Andreas Lochbihler topic = Computer science/Security/Cryptography date = 2019-10-07 notify = dbutler@turing.ac.uk abstract = We use CryptHOL to formalise commitment schemes and Sigma-protocols. Both are widely used fundamental two party cryptographic primitives. Security for commitment schemes is considered using game-based definitions whereas the security of Sigma-protocols is considered using both the game-based and simulation-based security paradigms. In this work, we first define security for both primitives and then prove secure multiple case studies: the Schnorr, Chaum-Pedersen and Okamoto Sigma-protocols as well as a construction that allows for compound (AND and OR statements) Sigma-protocols and the Pedersen and Rivest commitment schemes. We also prove that commitment schemes can be constructed from Sigma-protocols. We formalise this proof at an abstract level, only assuming the existence of a Sigma-protocol; consequently, the instantiations of this result for the concrete Sigma-protocols we consider come for free. [CryptHOL] title = CryptHOL author = Andreas Lochbihler topic = Computer science/Security/Cryptography, Computer science/Functional programming, Mathematics/Probability theory date = 2017-05-05 notify = mail@andreas-lochbihler.de abstract =

CryptHOL provides a framework for formalising cryptographic arguments in Isabelle/HOL. It shallowly embeds a probabilistic functional programming language in higher order logic. The language features monadic sequencing, recursion, random sampling, failures and failure handling, and black-box access to oracles. Oracles are probabilistic functions which maintain hidden state between different invocations. All operators are defined in the new semantic domain of generative probabilistic values, a codatatype. We derive proof rules for the operators and establish a connection with the theory of relational parametricity. Thus, the resuting proofs are trustworthy and comprehensible, and the framework is extensible and widely applicable.

The framework is used in the accompanying AFP entry "Game-based Cryptography in HOL". There, we show-case our framework by formalizing different game-based proofs from the literature. This formalisation continues the work described in the author's ESOP 2016 paper.

[Constructive_Cryptography] title = Constructive Cryptography in HOL author = Andreas Lochbihler , S. Reza Sefidgar<> topic = Computer science/Security/Cryptography, Mathematics/Probability theory date = 2018-12-17 notify = mail@andreas-lochbihler.de, reza.sefidgar@inf.ethz.ch abstract = Inspired by Abstract Cryptography, we extend CryptHOL, a framework for formalizing game-based proofs, with an abstract model of Random Systems and provide proof rules about their composition and equality. This foundation facilitates the formalization of Constructive Cryptography proofs, where the security of a cryptographic scheme is realized as a special form of construction in which a complex random system is built from simpler ones. This is a first step towards a fully-featured compositional framework, similar to Universal Composability framework, that supports formalization of simulation-based proofs. [Probabilistic_While] title = Probabilistic while loop author = Andreas Lochbihler topic = Computer science/Functional programming, Mathematics/Probability theory, Computer science/Algorithms date = 2017-05-05 notify = mail@andreas-lochbihler.de abstract = This AFP entry defines a probabilistic while operator based on sub-probability mass functions and formalises zero-one laws and variant rules for probabilistic loop termination. As applications, we implement probabilistic algorithms for the Bernoulli, geometric and arbitrary uniform distributions that only use fair coin flips, and prove them correct and terminating with probability 1. extra-history = Change history: [2018-02-02]: Added a proof that probabilistic conditioning can be implemented by repeated sampling. (revision 305867c4e911)
[Monad_Normalisation] title = Monad normalisation author = Joshua Schneider <>, Manuel Eberl , Andreas Lochbihler topic = Tools, Computer science/Functional programming, Logic/Rewriting date = 2017-05-05 notify = mail@andreas-lochbihler.de abstract = The usual monad laws can directly be used as rewrite rules for Isabelle’s simplifier to normalise monadic HOL terms and decide equivalences. In a commutative monad, however, the commutativity law is a higher-order permutative rewrite rule that makes the simplifier loop. This AFP entry implements a simproc that normalises monadic expressions in commutative monads using ordered rewriting. The simproc can also permute computations across control operators like if and case. [Monomorphic_Monad] title = Effect polymorphism in higher-order logic author = Andreas Lochbihler topic = Computer science/Functional programming date = 2017-05-05 notify = mail@andreas-lochbihler.de abstract = The notion of a monad cannot be expressed within higher-order logic (HOL) due to type system restrictions. We show that if a monad is used with values of only one type, this notion can be formalised in HOL. Based on this idea, we develop a library of effect specifications and implementations of monads and monad transformers. Hence, we can abstract over the concrete monad in HOL definitions and thus use the same definition for different (combinations of) effects. We illustrate the usefulness of effect polymorphism with a monadic interpreter for a simple language. extra-history = Change history: [2018-02-15]: added further specifications and implementations of non-determinism; more examples (revision bc5399eea78e)
[Constructor_Funs] title = Constructor Functions author = Lars Hupel topic = Tools date = 2017-04-19 notify = hupel@in.tum.de abstract = Isabelle's code generator performs various adaptations for target languages. Among others, constructor applications have to be fully saturated. That means that for constructor calls occuring as arguments to higher-order functions, synthetic lambdas have to be inserted. This entry provides tooling to avoid this construction altogether by introducing constructor functions. [Lazy_Case] title = Lazifying case constants author = Lars Hupel topic = Tools date = 2017-04-18 notify = hupel@in.tum.de abstract = Isabelle's code generator performs various adaptations for target languages. Among others, case statements are printed as match expressions. Internally, this is a sophisticated procedure, because in HOL, case statements are represented as nested calls to the case combinators as generated by the datatype package. Furthermore, the procedure relies on laziness of match expressions in the target language, i.e., that branches guarded by patterns that fail to match are not evaluated. Similarly, if-then-else is printed to the corresponding construct in the target language. This entry provides tooling to replace these special cases in the code generator by ignoring these target language features, instead printing case expressions and if-then-else as functions. [Dict_Construction] title = Dictionary Construction author = Lars Hupel topic = Tools date = 2017-05-24 notify = hupel@in.tum.de abstract = Isabelle's code generator natively supports type classes. For targets that do not have language support for classes and instances, it performs the well-known dictionary translation, as described by Haftmann and Nipkow. This translation happens outside the logic, i.e., there is no guarantee that it is correct, besides the pen-and-paper proof. This work implements a certified dictionary translation that produces new class-free constants and derives equality theorems. [Higher_Order_Terms] title = An Algebra for Higher-Order Terms author = Lars Hupel contributors = Yu Zhang <> topic = Computer science/Programming languages/Lambda calculi date = 2019-01-15 notify = lars@hupel.info abstract = In this formalization, I introduce a higher-order term algebra, generalizing the notions of free variables, matching, and substitution. The need arose from the work on a verified compiler from Isabelle to CakeML. Terms can be thought of as consisting of a generic (free variables, constants, application) and a specific part. As example applications, this entry provides instantiations for de-Bruijn terms, terms with named variables, and Blanchette’s λ-free higher-order terms. Furthermore, I implement translation functions between de-Bruijn terms and named terms and prove their correctness. [Subresultants] title = Subresultants author = Sebastiaan Joosten , René Thiemann , Akihisa Yamada topic = Mathematics/Algebra date = 2017-04-06 notify = rene.thiemann@uibk.ac.at abstract = We formalize the theory of subresultants and the subresultant polynomial remainder sequence as described by Brown and Traub. As a result, we obtain efficient certified algorithms for computing the resultant and the greatest common divisor of polynomials. [Comparison_Sort_Lower_Bound] title = Lower bound on comparison-based sorting algorithms author = Manuel Eberl topic = Computer science/Algorithms date = 2017-03-15 notify = manuel@pruvisto.org abstract =

This article contains a formal proof of the well-known fact that number of comparisons that a comparison-based sorting algorithm needs to perform to sort a list of length n is at least log2 (n!) in the worst case, i. e. Ω(n log n).

For this purpose, a shallow embedding for comparison-based sorting algorithms is defined: a sorting algorithm is a recursive datatype containing either a HOL function or a query of a comparison oracle with a continuation containing the remaining computation. This makes it possible to force the algorithm to use only comparisons and to track the number of comparisons made.

[Quick_Sort_Cost] title = The number of comparisons in QuickSort author = Manuel Eberl topic = Computer science/Algorithms date = 2017-03-15 notify = manuel@pruvisto.org abstract =

We give a formal proof of the well-known results about the number of comparisons performed by two variants of QuickSort: first, the expected number of comparisons of randomised QuickSort (i. e. QuickSort with random pivot choice) is 2 (n+1) Hn - 4 n, which is asymptotically equivalent to 2 n ln n; second, the number of comparisons performed by the classic non-randomised QuickSort has the same distribution in the average case as the randomised one.

[Random_BSTs] title = Expected Shape of Random Binary Search Trees author = Manuel Eberl topic = Computer science/Data structures date = 2017-04-04 notify = manuel@pruvisto.org abstract =

This entry contains proofs for the textbook results about the distributions of the height and internal path length of random binary search trees (BSTs), i. e. BSTs that are formed by taking an empty BST and inserting elements from a fixed set in random order.

In particular, we prove a logarithmic upper bound on the expected height and the Θ(n log n) closed-form solution for the expected internal path length in terms of the harmonic numbers. We also show how the internal path length relates to the average-case cost of a lookup in a BST.

[Randomised_BSTs] title = Randomised Binary Search Trees author = Manuel Eberl topic = Computer science/Data structures date = 2018-10-19 notify = manuel@pruvisto.org abstract =

This work is a formalisation of the Randomised Binary Search Trees introduced by Martínez and Roura, including definitions and correctness proofs.

Like randomised treaps, they are a probabilistic data structure that behaves exactly as if elements were inserted into a non-balancing BST in random order. However, unlike treaps, they only use discrete probability distributions, but their use of randomness is more complicated.

[E_Transcendental] title = The Transcendence of e author = Manuel Eberl topic = Mathematics/Analysis, Mathematics/Number theory date = 2017-01-12 notify = manuel@pruvisto.org abstract =

This work contains a proof that Euler's number e is transcendental. The proof follows the standard approach of assuming that e is algebraic and then using a specific integer polynomial to derive two inconsistent bounds, leading to a contradiction.

This kind of approach can be found in many different sources; this formalisation mostly follows a PlanetMath article by Roger Lipsett.

[Pi_Transcendental] title = The Transcendence of π author = Manuel Eberl topic = Mathematics/Number theory date = 2018-09-28 notify = manuel@pruvisto.org abstract =

This entry shows the transcendence of π based on the classic proof using the fundamental theorem of symmetric polynomials first given by von Lindemann in 1882, but the formalisation mostly follows the version by Niven. The proof reuses much of the machinery developed in the AFP entry on the transcendence of e.

[Hermite_Lindemann] title = The Hermite–Lindemann–Weierstraß Transcendence Theorem author = Manuel Eberl topic = Mathematics/Number theory date = 2021-03-03 notify = manuel@pruvisto.org abstract =

This article provides a formalisation of the Hermite-Lindemann-Weierstraß Theorem (also known as simply Hermite-Lindemann or Lindemann-Weierstraß). This theorem is one of the crowning achievements of 19th century number theory.

The theorem states that if $\alpha_1, \ldots, \alpha_n\in\mathbb{C}$ are algebraic numbers that are linearly independent over $\mathbb{Z}$, then $e^{\alpha_1},\ldots,e^{\alpha_n}$ are algebraically independent over $\mathbb{Q}$.

Like the previous formalisation in Coq by Bernard, I proceeded by formalising Baker's version of the theorem and proof and then deriving the original one from that. Baker's version states that for any algebraic numbers $\beta_1, \ldots, \beta_n\in\mathbb{C}$ and distinct algebraic numbers $\alpha_i, \ldots, \alpha_n\in\mathbb{C}$, we have $\beta_1 e^{\alpha_1} + \ldots + \beta_n e^{\alpha_n} = 0$ if and only if all the $\beta_i$ are zero.

This has a number of direct corollaries, e.g.:

  • $e$ and $\pi$ are transcendental
  • $e^z$, $\sin z$, $\tan z$, etc. are transcendental for algebraic $z\in\mathbb{C}\setminus\{0\}$
  • $\ln z$ is transcendental for algebraic $z\in\mathbb{C}\setminus\{0, 1\}$
[DFS_Framework] title = A Framework for Verifying Depth-First Search Algorithms author = Peter Lammich , René Neumann notify = lammich@in.tum.de date = 2016-07-05 topic = Computer science/Algorithms/Graph abstract =

This entry presents a framework for the modular verification of DFS-based algorithms, which is described in our [CPP-2015] paper. It provides a generic DFS algorithm framework, that can be parameterized with user-defined actions on certain events (e.g. discovery of new node). It comes with an extensible library of invariants, which can be used to derive invariants of a specific parameterization. Using refinement techniques, efficient implementations of the algorithms can easily be derived. Here, the framework comes with templates for a recursive and a tail-recursive implementation, and also with several templates for implementing the data structures required by the DFS algorithm. Finally, this entry contains a set of re-usable DFS-based algorithms, which illustrate the application of the framework.

[CPP-2015] Peter Lammich, René Neumann: A Framework for Verifying Depth-First Search Algorithms. CPP 2015: 137-146

[Flow_Networks] title = Flow Networks and the Min-Cut-Max-Flow Theorem author = Peter Lammich , S. Reza Sefidgar <> topic = Mathematics/Graph theory date = 2017-06-01 notify = lammich@in.tum.de abstract = We present a formalization of flow networks and the Min-Cut-Max-Flow theorem. Our formal proof closely follows a standard textbook proof, and is accessible even without being an expert in Isabelle/HOL, the interactive theorem prover used for the formalization. [Prpu_Maxflow] title = Formalizing Push-Relabel Algorithms author = Peter Lammich , S. Reza Sefidgar <> topic = Computer science/Algorithms/Graph, Mathematics/Graph theory date = 2017-06-01 notify = lammich@in.tum.de abstract = We present a formalization of push-relabel algorithms for computing the maximum flow in a network. We start with Goldberg's et al.~generic push-relabel algorithm, for which we show correctness and the time complexity bound of O(V^2E). We then derive the relabel-to-front and FIFO implementation. Using stepwise refinement techniques, we derive an efficient verified implementation. Our formal proof of the abstract algorithms closely follows a standard textbook proof. It is accessible even without being an expert in Isabelle/HOL, the interactive theorem prover used for the formalization. [Buildings] title = Chamber Complexes, Coxeter Systems, and Buildings author = Jeremy Sylvestre notify = jeremy.sylvestre@ualberta.ca date = 2016-07-01 topic = Mathematics/Algebra, Mathematics/Geometry abstract = We provide a basic formal framework for the theory of chamber complexes and Coxeter systems, and for buildings as thick chamber complexes endowed with a system of apartments. Along the way, we develop some of the general theory of abstract simplicial complexes and of groups (relying on the group_add class for the basics), including free groups and group presentations, and their universal properties. The main results verified are that the deletion condition is both necessary and sufficient for a group with a set of generators of order two to be a Coxeter system, and that the apartments in a (thick) building are all uniformly Coxeter. [Algebraic_VCs] title = Program Construction and Verification Components Based on Kleene Algebra author = Victor B. F. Gomes , Georg Struth notify = victor.gomes@cl.cam.ac.uk, g.struth@sheffield.ac.uk date = 2016-06-18 topic = Mathematics/Algebra abstract = Variants of Kleene algebra support program construction and verification by algebraic reasoning. This entry provides a verification component for Hoare logic based on Kleene algebra with tests, verification components for weakest preconditions and strongest postconditions based on Kleene algebra with domain and a component for step-wise refinement based on refinement Kleene algebra with tests. In addition to these components for the partial correctness of while programs, a verification component for total correctness based on divergence Kleene algebras and one for (partial correctness) of recursive programs based on domain quantales are provided. Finally we have integrated memory models for programs with pointers and a program trace semantics into the weakest precondition component. [C2KA_DistributedSystems] title = Communicating Concurrent Kleene Algebra for Distributed Systems Specification author = Maxime Buyse , Jason Jaskolka topic = Computer science/Automata and formal languages, Mathematics/Algebra date = 2019-08-06 notify = maxime.buyse@polytechnique.edu, jason.jaskolka@carleton.ca abstract = Communicating Concurrent Kleene Algebra (C²KA) is a mathematical framework for capturing the communicating and concurrent behaviour of agents in distributed systems. It extends Hoare et al.'s Concurrent Kleene Algebra (CKA) with communication actions through the notions of stimuli and shared environments. C²KA has applications in studying system-level properties of distributed systems such as safety, security, and reliability. In this work, we formalize results about C²KA and its application for distributed systems specification. We first formalize the stimulus structure and behaviour structure (CKA). Next, we combine them to formalize C²KA and its properties. Then, we formalize notions and properties related to the topology of distributed systems and the potential for communication via stimuli and via shared environments of agents, all within the algebraic setting of C²KA. [Card_Equiv_Relations] title = Cardinality of Equivalence Relations author = Lukas Bulwahn notify = lukas.bulwahn@gmail.com date = 2016-05-24 topic = Mathematics/Combinatorics abstract = This entry provides formulae for counting the number of equivalence relations and partial equivalence relations over a finite carrier set with given cardinality. To count the number of equivalence relations, we provide bijections between equivalence relations and set partitions, and then transfer the main results of the two AFP entries, Cardinality of Set Partitions and Spivey's Generalized Recurrence for Bell Numbers, to theorems on equivalence relations. To count the number of partial equivalence relations, we observe that counting partial equivalence relations over a set A is equivalent to counting all equivalence relations over all subsets of the set A. From this observation and the results on equivalence relations, we show that the cardinality of partial equivalence relations over a finite set of cardinality n is equal to the n+1-th Bell number. [Twelvefold_Way] title = The Twelvefold Way author = Lukas Bulwahn topic = Mathematics/Combinatorics date = 2016-12-29 notify = lukas.bulwahn@gmail.com abstract = This entry provides all cardinality theorems of the Twelvefold Way. The Twelvefold Way systematically classifies twelve related combinatorial problems concerning two finite sets, which include counting permutations, combinations, multisets, set partitions and number partitions. This development builds upon the existing formal developments with cardinality theorems for those structures. It provides twelve bijections from the various structures to different equivalence classes on finite functions, and hence, proves cardinality formulae for these equivalence classes on finite functions. [Chord_Segments] title = Intersecting Chords Theorem author = Lukas Bulwahn notify = lukas.bulwahn@gmail.com date = 2016-10-11 topic = Mathematics/Geometry abstract = This entry provides a geometric proof of the intersecting chords theorem. The theorem states that when two chords intersect each other inside a circle, the products of their segments are equal. After a short review of existing proofs in the literature, I decided to use a proof approach that employs reasoning about lengths of line segments, the orthogonality of two lines and the Pythagoras Law. Hence, one can understand the formalized proof easily with the knowledge of a few general geometric facts that are commonly taught in high-school. This theorem is the 55th theorem of the Top 100 Theorems list. [Category3] title = Category Theory with Adjunctions and Limits author = Eugene W. Stark notify = stark@cs.stonybrook.edu date = 2016-06-26 topic = Mathematics/Category theory abstract =

This article attempts to develop a usable framework for doing category theory in Isabelle/HOL. Our point of view, which to some extent differs from that of the previous AFP articles on the subject, is to try to explore how category theory can be done efficaciously within HOL, rather than trying to match exactly the way things are done using a traditional approach. To this end, we define the notion of category in an "object-free" style, in which a category is represented by a single partial composition operation on arrows. This way of defining categories provides some advantages in the context of HOL, including the ability to avoid the use of records and the possibility of defining functors and natural transformations simply as certain functions on arrows, rather than as composite objects. We define various constructions associated with the basic notions, including: dual category, product category, functor category, discrete category, free category, functor composition, and horizontal and vertical composite of natural transformations. A "set category" locale is defined that axiomatizes the notion "category of all sets at a type and all functions between them," and a fairly extensive set of properties of set categories is derived from the locale assumptions. The notion of a set category is used to prove the Yoneda Lemma in a general setting of a category equipped with a "hom embedding," which maps arrows of the category to the "universe" of the set category. We also give a treatment of adjunctions, defining adjunctions via left and right adjoint functors, natural bijections between hom-sets, and unit and counit natural transformations, and showing the equivalence of these definitions. We also develop the theory of limits, including representations of functors, diagrams and cones, and diagonal functors. We show that right adjoint functors preserve limits, and that limits can be constructed via products and equalizers. We characterize the conditions under which limits exist in a set category. We also examine the case of limits in a functor category, ultimately culminating in a proof that the Yoneda embedding preserves limits.

Revisions made subsequent to the first version of this article added material on equivalence of categories, cartesian categories, categories with pullbacks, categories with finite limits, and cartesian closed categories. A construction was given of the category of hereditarily finite sets and functions between them, and it was shown that this category is cartesian closed.

extra-history = Change history: [2018-05-29]: Revised axioms for the category locale. Introduced notation for composition and "in hom". (revision 8318366d4575)
[2020-02-15]: Move ConcreteCategory.thy from Bicategory to Category3 and use it systematically. Make other minor improvements throughout. (revision a51840d36867)
[2020-07-10]: Added new material, mostly centered around cartesian categories. (revision 06640f317a79)
[2020-11-04]: Minor modifications and extensions made in conjunction with the addition of new material to Bicategory. (revision 472cb2268826)
[2021-07-22]: Minor changes to sublocale declarations related to functor/natural transformation to avoid issues with global interpretations reported 2/2/2021 by Filip Smola. (revision 49d3aa43c180)
[MonoidalCategory] title = Monoidal Categories author = Eugene W. Stark topic = Mathematics/Category theory date = 2017-05-04 notify = stark@cs.stonybrook.edu abstract =

Building on the formalization of basic category theory set out in the author's previous AFP article, the present article formalizes some basic aspects of the theory of monoidal categories. Among the notions defined here are monoidal category, monoidal functor, and equivalence of monoidal categories. The main theorems formalized are MacLane's coherence theorem and the constructions of the free monoidal category and free strict monoidal category generated by a given category. The coherence theorem is proved syntactically, using a structurally recursive approach to reduction of terms that might have some novel aspects. We also give proofs of some results given by Etingof et al, which may prove useful in a formal setting. In particular, we show that the left and right unitors need not be taken as given data in the definition of monoidal category, nor does the definition of monoidal functor need to take as given a specific isomorphism expressing the preservation of the unit object. Our definitions of monoidal category and monoidal functor are stated so as to take advantage of the economy afforded by these facts.

Revisions made subsequent to the first version of this article added material on cartesian monoidal categories; showing that the underlying category of a cartesian monoidal category is a cartesian category, and that every cartesian category extends to a cartesian monoidal category.

extra-history = Change history: [2017-05-18]: Integrated material from MonoidalCategory/Category3Adapter into Category3/ and deleted adapter. (revision 015543cdd069)
[2018-05-29]: Modifications required due to 'Category3' changes. Introduced notation for "in hom". (revision 8318366d4575)
[2020-02-15]: Cosmetic improvements. (revision a51840d36867)
[2020-07-10]: Added new material on cartesian monoidal categories. (revision 06640f317a79)
[Card_Multisets] title = Cardinality of Multisets author = Lukas Bulwahn notify = lukas.bulwahn@gmail.com date = 2016-06-26 topic = Mathematics/Combinatorics abstract =

This entry provides three lemmas to count the number of multisets of a given size and finite carrier set. The first lemma provides a cardinality formula assuming that the multiset's elements are chosen from the given carrier set. The latter two lemmas provide formulas assuming that the multiset's elements also cover the given carrier set, i.e., each element of the carrier set occurs in the multiset at least once.

The proof of the first lemma uses the argument of the recurrence relation for counting multisets. The proof of the second lemma is straightforward, and the proof of the third lemma is easily obtained using the first cardinality lemma. A challenge for the formalization is the derivation of the required induction rule, which is a special combination of the induction rules for finite sets and natural numbers. The induction rule is derived by defining a suitable inductive predicate and transforming the predicate's induction rule.

[Posix-Lexing] title = POSIX Lexing with Derivatives of Regular Expressions author = Fahad Ausaf , Roy Dyckhoff , Christian Urban notify = christian.urban@kcl.ac.uk date = 2016-05-24 topic = Computer science/Automata and formal languages abstract = Brzozowski introduced the notion of derivatives for regular expressions. They can be used for a very simple regular expression matching algorithm. Sulzmann and Lu cleverly extended this algorithm in order to deal with POSIX matching, which is the underlying disambiguation strategy for regular expressions needed in lexers. In this entry we give our inductive definition of what a POSIX value is and show (i) that such a value is unique (for given regular expression and string being matched) and (ii) that Sulzmann and Lu's algorithm always generates such a value (provided that the regular expression matches the string). We also prove the correctness of an optimised version of the POSIX matching algorithm. [LocalLexing] title = Local Lexing author = Steven Obua topic = Computer science/Automata and formal languages date = 2017-04-28 notify = steven@recursivemind.com abstract = This formalisation accompanies the paper Local Lexing which introduces a novel parsing concept of the same name. The paper also gives a high-level algorithm for local lexing as an extension of Earley's algorithm. This formalisation proves the algorithm to be correct with respect to its local lexing semantics. As a special case, this formalisation thus also contains a proof of the correctness of Earley's algorithm. The paper contains a short outline of how this formalisation is organised. [MFMC_Countable] title = A Formal Proof of the Max-Flow Min-Cut Theorem for Countable Networks author = Andreas Lochbihler date = 2016-05-09 topic = Mathematics/Graph theory abstract = This article formalises a proof of the maximum-flow minimal-cut theorem for networks with countably many edges. A network is a directed graph with non-negative real-valued edge labels and two dedicated vertices, the source and the sink. A flow in a network assigns non-negative real numbers to the edges such that for all vertices except for the source and the sink, the sum of values on incoming edges equals the sum of values on outgoing edges. A cut is a subset of the vertices which contains the source, but not the sink. Our theorem states that in every network, there is a flow and a cut such that the flow saturates all the edges going out of the cut and is zero on all the incoming edges. The proof is based on the paper The Max-Flow Min-Cut theorem for countable networks by Aharoni et al. Additionally, we prove a characterisation of the lifting operation for relations on discrete probability distributions, which leads to a concise proof of its distributivity over relation composition. notify = mail@andreas-lochbihler.de extra-history = Change history: [2017-09-06]: derive characterisation for the lifting operation on discrete distributions from finite version of the max-flow min-cut theorem (revision a7a198f5bab0)
[2020-12-19]: simpler proof of linkability for bounded unhindered bipartite webs, leading to a simpler proof for networks with bounded out-capacities (revision 93ca33f4d915)
[2021-08-13]: generalize the derivation of the characterisation for the relator of discrete probability distributions to work for the bounded and unbounded MFMC theorem (revision 3c85bb52bbe6)
[Liouville_Numbers] title = Liouville numbers author = Manuel Eberl date = 2015-12-28 topic = Mathematics/Analysis, Mathematics/Number theory abstract =

Liouville numbers are a class of transcendental numbers that can be approximated particularly well with rational numbers. Historically, they were the first numbers whose transcendence was proven.

In this entry, we define the concept of Liouville numbers as well as the standard construction to obtain Liouville numbers (including Liouville's constant) and we prove their most important properties: irrationality and transcendence.

The proof is very elementary and requires only standard arithmetic, the Mean Value Theorem for polynomials, and the boundedness of polynomials on compact intervals.

notify = manuel@pruvisto.org [Triangle] title = Basic Geometric Properties of Triangles author = Manuel Eberl date = 2015-12-28 topic = Mathematics/Geometry abstract =

This entry contains a definition of angles between vectors and between three points. Building on this, we prove basic geometric properties of triangles, such as the Isosceles Triangle Theorem, the Law of Sines and the Law of Cosines, that the sum of the angles of a triangle is π, and the congruence theorems for triangles.

The definitions and proofs were developed following those by John Harrison in HOL Light. However, due to Isabelle's type class system, all definitions and theorems in the Isabelle formalisation hold for all real inner product spaces.

notify = manuel@pruvisto.org [Prime_Harmonic_Series] title = The Divergence of the Prime Harmonic Series author = Manuel Eberl date = 2015-12-28 topic = Mathematics/Number theory abstract =

In this work, we prove the lower bound ln(H_n) - ln(5/3) for the partial sum of the Prime Harmonic series and, based on this, the divergence of the Prime Harmonic Series ∑[p prime] · 1/p.

The proof relies on the unique squarefree decomposition of natural numbers. This is similar to Euler's original proof (which was highly informal and morally questionable). Its advantage over proofs by contradiction, like the famous one by Paul Erdős, is that it provides a relatively good lower bound for the partial sums.

notify = manuel@pruvisto.org [Descartes_Sign_Rule] title = Descartes' Rule of Signs author = Manuel Eberl date = 2015-12-28 topic = Mathematics/Analysis abstract =

Descartes' Rule of Signs relates the number of positive real roots of a polynomial with the number of sign changes in its coefficient sequence.

Our proof follows the simple inductive proof given by Rob Arthan, which was also used by John Harrison in his HOL Light formalisation. We proved most of the lemmas for arbitrary linearly-ordered integrity domains (e.g. integers, rationals, reals); the main result, however, requires the intermediate value theorem and was therefore only proven for real polynomials.

notify = manuel@pruvisto.org [Euler_MacLaurin] title = The Euler–MacLaurin Formula author = Manuel Eberl topic = Mathematics/Analysis date = 2017-03-10 notify = manuel@pruvisto.org abstract =

The Euler-MacLaurin formula relates the value of a discrete sum to that of the corresponding integral in terms of the derivatives at the borders of the summation and a remainder term. Since the remainder term is often very small as the summation bounds grow, this can be used to compute asymptotic expansions for sums.

This entry contains a proof of this formula for functions from the reals to an arbitrary Banach space. Two variants of the formula are given: the standard textbook version and a variant outlined in Concrete Mathematics that is more useful for deriving asymptotic estimates.

As example applications, we use that formula to derive the full asymptotic expansion of the harmonic numbers and the sum of inverse squares.

[Card_Partitions] title = Cardinality of Set Partitions author = Lukas Bulwahn date = 2015-12-12 topic = Mathematics/Combinatorics abstract = The theory's main theorem states that the cardinality of set partitions of size k on a carrier set of size n is expressed by Stirling numbers of the second kind. In Isabelle, Stirling numbers of the second kind are defined in the AFP entry `Discrete Summation` through their well-known recurrence relation. The main theorem relates them to the alternative definition as cardinality of set partitions. The proof follows the simple and short explanation in Richard P. Stanley's `Enumerative Combinatorics: Volume 1` and Wikipedia, and unravels the full details and implicit reasoning steps of these explanations. notify = lukas.bulwahn@gmail.com [Card_Number_Partitions] title = Cardinality of Number Partitions author = Lukas Bulwahn date = 2016-01-14 topic = Mathematics/Combinatorics abstract = This entry provides a basic library for number partitions, defines the two-argument partition function through its recurrence relation and relates this partition function to the cardinality of number partitions. The main proof shows that the recursively-defined partition function with arguments n and k equals the cardinality of number partitions of n with exactly k parts. The combinatorial proof follows the proof sketch of Theorem 2.4.1 in Mazur's textbook `Combinatorics: A Guided Tour`. This entry can serve as starting point for various more intrinsic properties about number partitions, the partition function and related recurrence relations. notify = lukas.bulwahn@gmail.com [Multirelations] title = Binary Multirelations author = Hitoshi Furusawa , Georg Struth date = 2015-06-11 topic = Mathematics/Algebra abstract = Binary multirelations associate elements of a set with its subsets; hence they are binary relations from a set to its power set. Applications include alternating automata, models and logics for games, program semantics with dual demonic and angelic nondeterministic choices and concurrent dynamic logics. This proof document supports an arXiv article that formalises the basic algebra of multirelations and proposes axiom systems for them, ranging from weak bi-monoids to weak bi-quantales. notify = [Noninterference_Generic_Unwinding] title = The Generic Unwinding Theorem for CSP Noninterference Security author = Pasquale Noce date = 2015-06-11 topic = Computer science/Security, Computer science/Concurrency/Process calculi abstract =

The classical definition of noninterference security for a deterministic state machine with outputs requires to consider the outputs produced by machine actions after any trace, i.e. any indefinitely long sequence of actions, of the machine. In order to render the verification of the security of such a machine more straightforward, there is a need of some sufficient condition for security such that just individual actions, rather than unbounded sequences of actions, have to be considered.

By extending previous results applying to transitive noninterference policies, Rushby has proven an unwinding theorem that provides a sufficient condition of this kind in the general case of a possibly intransitive policy. This condition has to be satisfied by a generic function mapping security domains into equivalence relations over machine states.

An analogous problem arises for CSP noninterference security, whose definition requires to consider any possible future, i.e. any indefinitely long sequence of subsequent events and any indefinitely large set of refused events associated to that sequence, for each process trace.

This paper provides a sufficient condition for CSP noninterference security, which indeed requires to just consider individual accepted and refused events and applies to the general case of a possibly intransitive policy. This condition follows Rushby's one for classical noninterference security, and has to be satisfied by a generic function mapping security domains into equivalence relations over process traces; hence its name, Generic Unwinding Theorem. Variants of this theorem applying to deterministic processes and trace set processes are also proven. Finally, the sufficient condition for security expressed by the theorem is shown not to be a necessary condition as well, viz. there exists a secure process such that no domain-relation map satisfying the condition exists.

notify = [Noninterference_Ipurge_Unwinding] title = The Ipurge Unwinding Theorem for CSP Noninterference Security author = Pasquale Noce date = 2015-06-11 topic = Computer science/Security abstract =

The definition of noninterference security for Communicating Sequential Processes requires to consider any possible future, i.e. any indefinitely long sequence of subsequent events and any indefinitely large set of refused events associated to that sequence, for each process trace. In order to render the verification of the security of a process more straightforward, there is a need of some sufficient condition for security such that just individual accepted and refused events, rather than unbounded sequences and sets of events, have to be considered.

Of course, if such a sufficient condition were necessary as well, it would be even more valuable, since it would permit to prove not only that a process is secure by verifying that the condition holds, but also that a process is not secure by verifying that the condition fails to hold.

This paper provides a necessary and sufficient condition for CSP noninterference security, which indeed requires to just consider individual accepted and refused events and applies to the general case of a possibly intransitive policy. This condition follows Rushby's output consistency for deterministic state machines with outputs, and has to be satisfied by a specific function mapping security domains into equivalence relations over process traces. The definition of this function makes use of an intransitive purge function following Rushby's one; hence the name given to the condition, Ipurge Unwinding Theorem.

Furthermore, in accordance with Hoare's formal definition of deterministic processes, it is shown that a process is deterministic just in case it is a trace set process, i.e. it may be identified by means of a trace set alone, matching the set of its traces, in place of a failures-divergences pair. Then, variants of the Ipurge Unwinding Theorem are proven for deterministic processes and trace set processes.

notify = [Relational_Method] title = The Relational Method with Message Anonymity for the Verification of Cryptographic Protocols author = Pasquale Noce topic = Computer science/Security date = 2020-12-05 notify = pasquale.noce.lavoro@gmail.com abstract = This paper introduces a new method for the formal verification of cryptographic protocols, the relational method, derived from Paulson's inductive method by means of some enhancements aimed at streamlining formal definitions and proofs, specially for protocols using public key cryptography. Moreover, this paper proposes a method to formalize a further security property, message anonymity, in addition to message confidentiality and authenticity. The relational method, including message anonymity, is then applied to the verification of a sample authentication protocol, comprising Password Authenticated Connection Establishment (PACE) with Chip Authentication Mapping followed by the explicit verification of an additional password over the PACE secure channel. [List_Interleaving] title = Reasoning about Lists via List Interleaving author = Pasquale Noce date = 2015-06-11 topic = Computer science/Data structures abstract =

Among the various mathematical tools introduced in his outstanding work on Communicating Sequential Processes, Hoare has defined "interleaves" as the predicate satisfied by any three lists such that the first list may be split into sublists alternately extracted from the other two ones, whatever is the criterion for extracting an item from either one list or the other in each step.

This paper enriches Hoare's definition by identifying such criterion with the truth value of a predicate taking as inputs the head and the tail of the first list. This enhanced "interleaves" predicate turns out to permit the proof of equalities between lists without the need of an induction. Some rules that allow to infer "interleaves" statements without induction, particularly applying to the addition or removal of a prefix to the input lists, are also proven. Finally, a stronger version of the predicate, named "Interleaves", is shown to fulfil further rules applying to the addition or removal of a suffix to the input lists.

notify = [Residuated_Lattices] title = Residuated Lattices author = Victor B. F. Gomes , Georg Struth date = 2015-04-15 topic = Mathematics/Algebra abstract = The theory of residuated lattices, first proposed by Ward and Dilworth, is formalised in Isabelle/HOL. This includes concepts of residuated functions; their adjoints and conjugates. It also contains necessary and sufficient conditions for the existence of these operations in an arbitrary lattice. The mathematical components for residuated lattices are linked to the AFP entry for relation algebra. In particular, we prove Jonsson and Tsinakis conditions for a residuated boolean algebra to form a relation algebra. notify = g.struth@sheffield.ac.uk [ConcurrentGC] title = Relaxing Safely: Verified On-the-Fly Garbage Collection for x86-TSO author = Peter Gammie , Tony Hosking , Kai Engelhardt <> date = 2015-04-13 topic = Computer science/Algorithms/Concurrent abstract =

We use ConcurrentIMP to model Schism, a state-of-the-art real-time garbage collection scheme for weak memory, and show that it is safe on x86-TSO.

This development accompanies the PLDI 2015 paper of the same name.

notify = peteg42@gmail.com [List_Update] title = Analysis of List Update Algorithms author = Maximilian P.L. Haslbeck , Tobias Nipkow date = 2016-02-17 topic = Computer science/Algorithms/Online abstract =

These theories formalize the quantitative analysis of a number of classical algorithms for the list update problem: 2-competitiveness of move-to-front, the lower bound of 2 for the competitiveness of deterministic list update algorithms and 1.6-competitiveness of the randomized COMB algorithm, the best randomized list update algorithm known to date. The material is based on the first two chapters of Online Computation and Competitive Analysis by Borodin and El-Yaniv.

For an informal description see the FSTTCS 2016 publication Verified Analysis of List Update Algorithms by Haslbeck and Nipkow.

notify = nipkow@in.tum.de [ConcurrentIMP] title = Concurrent IMP author = Peter Gammie date = 2015-04-13 topic = Computer science/Programming languages/Logics abstract = ConcurrentIMP extends the small imperative language IMP with control non-determinism and constructs for synchronous message passing. notify = peteg42@gmail.com [TortoiseHare] title = The Tortoise and Hare Algorithm author = Peter Gammie date = 2015-11-18 topic = Computer science/Algorithms abstract = We formalize the Tortoise and Hare cycle-finding algorithm ascribed to Floyd by Knuth, and an improved version due to Brent. notify = peteg42@gmail.com [UPF] title = The Unified Policy Framework (UPF) author = Achim D. Brucker , Lukas Brügger , Burkhart Wolff date = 2014-11-28 topic = Computer science/Security abstract = We present the Unified Policy Framework (UPF), a generic framework for modelling security (access-control) policies. UPF emphasizes the view that a policy is a policy decision function that grants or denies access to resources, permissions, etc. In other words, instead of modelling the relations of permitted or prohibited requests directly, we model the concrete function that implements the policy decision point in a system. In more detail, UPF is based on the following four principles: 1) Functional representation of policies, 2) No conflicts are possible, 3) Three-valued decision type (allow, deny, undefined), 4) Output type not containing the decision only. notify = adbrucker@0x5f.org, wolff@lri.fr, lukas.a.bruegger@gmail.com [UPF_Firewall] title = Formal Network Models and Their Application to Firewall Policies author = Achim D. Brucker , Lukas Brügger<>, Burkhart Wolff topic = Computer science/Security, Computer science/Networks date = 2017-01-08 notify = adbrucker@0x5f.org abstract = We present a formal model of network protocols and their application to modeling firewall policies. The formalization is based on the Unified Policy Framework (UPF). The formalization was originally developed with for generating test cases for testing the security configuration actual firewall and router (middle-boxes) using HOL-TestGen. Our work focuses on modeling application level protocols on top of tcp/ip. [AODV] title = Loop freedom of the (untimed) AODV routing protocol author = Timothy Bourke , Peter Höfner date = 2014-10-23 topic = Computer science/Concurrency/Process calculi abstract =

The Ad hoc On-demand Distance Vector (AODV) routing protocol allows the nodes in a Mobile Ad hoc Network (MANET) or a Wireless Mesh Network (WMN) to know where to forward data packets. Such a protocol is ‘loop free’ if it never leads to routing decisions that forward packets in circles.

This development mechanises an existing pen-and-paper proof of loop freedom of AODV. The protocol is modelled in the Algebra of Wireless Networks (AWN), which is the subject of an earlier paper and AFP mechanization. The proof relies on a novel compositional approach for lifting invariants to networks of nodes.

We exploit the mechanization to analyse several variants of AODV and show that Isabelle/HOL can re-establish most proof obligations automatically and identify exactly the steps that are no longer valid.

notify = tim@tbrk.org [Show] title = Haskell's Show Class in Isabelle/HOL author = Christian Sternagel , René Thiemann date = 2014-07-29 topic = Computer science/Functional programming license = LGPL abstract = We implemented a type class for "to-string" functions, similar to Haskell's Show class. Moreover, we provide instantiations for Isabelle/HOL's standard types like bool, prod, sum, nats, ints, and rats. It is further possible, to automatically derive show functions for arbitrary user defined datatypes similar to Haskell's "deriving Show". extra-history = Change history: [2015-03-11]: Adapted development to new-style (BNF-based) datatypes.
[2015-04-10]: Moved development for old-style datatypes into subdirectory "Old_Datatype".
notify = christian.sternagel@uibk.ac.at, rene.thiemann@uibk.ac.at [Certification_Monads] title = Certification Monads author = Christian Sternagel , René Thiemann date = 2014-10-03 topic = Computer science/Functional programming abstract = This entry provides several monads intended for the development of stand-alone certifiers via code generation from Isabelle/HOL. More specifically, there are three flavors of error monads (the sum type, for the case where all monadic functions are total; an instance of the former, the so called check monad, yielding either success without any further information or an error message; as well as a variant of the sum type that accommodates partial functions by providing an explicit bottom element) and a parser monad built on top. All of this monads are heavily used in the IsaFoR/CeTA project which thus provides many examples of their usage. notify = c.sternagel@gmail.com, rene.thiemann@uibk.ac.at [CISC-Kernel] title = Formal Specification of a Generic Separation Kernel author = Freek Verbeek , Sergey Tverdyshev , Oto Havle , Holger Blasum , Bruno Langenstein , Werner Stephan , Yakoub Nemouchi , Abderrahmane Feliachi , Burkhart Wolff , Julien Schmaltz date = 2014-07-18 topic = Computer science/Security abstract =

Intransitive noninterference has been a widely studied topic in the last few decades. Several well-established methodologies apply interactive theorem proving to formulate a noninterference theorem over abstract academic models. In joint work with several industrial and academic partners throughout Europe, we are helping in the certification process of PikeOS, an industrial separation kernel developed at SYSGO. In this process, established theories could not be applied. We present a new generic model of separation kernels and a new theory of intransitive noninterference. The model is rich in detail, making it suitable for formal verification of realistic and industrial systems such as PikeOS. Using a refinement-based theorem proving approach, we ensure that proofs remain manageable.

This document corresponds to the deliverable D31.1 of the EURO-MILS Project http://www.euromils.eu.

notify = [pGCL] title = pGCL for Isabelle author = David Cock date = 2014-07-13 topic = Computer science/Programming languages/Language definitions abstract =

pGCL is both a programming language and a specification language that incorporates both probabilistic and nondeterministic choice, in a unified manner. Program verification is by refinement or annotation (or both), using either Hoare triples, or weakest-precondition entailment, in the style of GCL.

This package provides both a shallow embedding of the language primitives, and an annotation and refinement framework. The generated document includes a brief tutorial.

notify = [Noninterference_CSP] title = Noninterference Security in Communicating Sequential Processes author = Pasquale Noce date = 2014-05-23 topic = Computer science/Security abstract =

An extension of classical noninterference security for deterministic state machines, as introduced by Goguen and Meseguer and elegantly formalized by Rushby, to nondeterministic systems should satisfy two fundamental requirements: it should be based on a mathematically precise theory of nondeterminism, and should be equivalent to (or at least not weaker than) the classical notion in the degenerate deterministic case.

This paper proposes a definition of noninterference security applying to Hoare's Communicating Sequential Processes (CSP) in the general case of a possibly intransitive noninterference policy, and proves the equivalence of this security property to classical noninterference security for processes representing deterministic state machines.

Furthermore, McCullough's generalized noninterference security is shown to be weaker than both the proposed notion of CSP noninterference security for a generic process, and classical noninterference security for processes representing deterministic state machines. This renders CSP noninterference security preferable as an extension of classical noninterference security to nondeterministic systems.

notify = pasquale.noce.lavoro@gmail.com [Floyd_Warshall] title = The Floyd-Warshall Algorithm for Shortest Paths author = Simon Wimmer , Peter Lammich topic = Computer science/Algorithms/Graph date = 2017-05-08 notify = wimmers@in.tum.de abstract = The Floyd-Warshall algorithm [Flo62, Roy59, War62] is a classic dynamic programming algorithm to compute the length of all shortest paths between any two vertices in a graph (i.e. to solve the all-pairs shortest path problem, or APSP for short). Given a representation of the graph as a matrix of weights M, it computes another matrix M' which represents a graph with the same path lengths and contains the length of the shortest path between any two vertices i and j. This is only possible if the graph does not contain any negative cycles. However, in this case the Floyd-Warshall algorithm will detect the situation by calculating a negative diagonal entry. This entry includes a formalization of the algorithm and of these key properties. The algorithm is refined to an efficient imperative version using the Imperative Refinement Framework. [Roy_Floyd_Warshall] title = Transitive closure according to Roy-Floyd-Warshall author = Makarius Wenzel <> date = 2014-05-23 topic = Computer science/Algorithms/Graph abstract = This formulation of the Roy-Floyd-Warshall algorithm for the transitive closure bypasses matrices and arrays, but uses a more direct mathematical model with adjacency functions for immediate predecessors and successors. This can be implemented efficiently in functional programming languages and is particularly adequate for sparse relations. notify = [GPU_Kernel_PL] title = Syntax and semantics of a GPU kernel programming language author = John Wickerson date = 2014-04-03 topic = Computer science/Programming languages/Language definitions abstract = This document accompanies the article "The Design and Implementation of a Verification Technique for GPU Kernels" by Adam Betts, Nathan Chong, Alastair F. Donaldson, Jeroen Ketema, Shaz Qadeer, Paul Thomson and John Wickerson. It formalises all of the definitions provided in Sections 3 and 4 of the article. notify = [AWN] title = Mechanization of the Algebra for Wireless Networks (AWN) author = Timothy Bourke date = 2014-03-08 topic = Computer science/Concurrency/Process calculi abstract =

AWN is a process algebra developed for modelling and analysing protocols for Mobile Ad hoc Networks (MANETs) and Wireless Mesh Networks (WMNs). AWN models comprise five distinct layers: sequential processes, local parallel compositions, nodes, partial networks, and complete networks.

This development mechanises the original operational semantics of AWN and introduces a variant 'open' operational semantics that enables the compositional statement and proof of invariants across distinct network nodes. It supports labels (for weakening invariants) and (abstract) data state manipulations. A framework for compositional invariant proofs is developed, including a tactic (inv_cterms) for inductive invariant proofs of sequential processes, lifting rules for the open versions of the higher layers, and a rule for transferring lifted properties back to the standard semantics. A notion of 'control terms' reduces proof obligations to the subset of subterms that act directly (in contrast to operators for combining terms and joining processes).

notify = tim@tbrk.org [Selection_Heap_Sort] title = Verification of Selection and Heap Sort Using Locales author = Danijela Petrovic date = 2014-02-11 topic = Computer science/Algorithms abstract = Stepwise program refinement techniques can be used to simplify program verification. Programs are better understood since their main properties are clearly stated, and verification of rather complex algorithms is reduced to proving simple statements connecting successive program specifications. Additionally, it is easy to analyze similar algorithms and to compare their properties within a single formalization. Usually, formal analysis is not done in educational setting due to complexity of verification and a lack of tools and procedures to make comparison easy. Verification of an algorithm should not only give correctness proof, but also better understanding of an algorithm. If the verification is based on small step program refinement, it can become simple enough to be demonstrated within the university-level computer science curriculum. In this paper we demonstrate this and give a formal analysis of two well known algorithms (Selection Sort and Heap Sort) using proof assistant Isabelle/HOL and program refinement techniques. notify = [Real_Impl] title = Implementing field extensions of the form Q[sqrt(b)] author = René Thiemann date = 2014-02-06 license = LGPL topic = Mathematics/Analysis abstract = We apply data refinement to implement the real numbers, where we support all numbers in the field extension Q[sqrt(b)], i.e., all numbers of the form p + q * sqrt(b) for rational numbers p and q and some fixed natural number b. To this end, we also developed algorithms to precisely compute roots of a rational number, and to perform a factorization of natural numbers which eliminates duplicate prime factors.

Our results have been used to certify termination proofs which involve polynomial interpretations over the reals. extra-history = Change history: [2014-07-11]: Moved NthRoot_Impl to Sqrt-Babylonian. notify = rene.thiemann@uibk.ac.at [ShortestPath] title = An Axiomatic Characterization of the Single-Source Shortest Path Problem author = Christine Rizkallah date = 2013-05-22 topic = Mathematics/Graph theory abstract = This theory is split into two sections. In the first section, we give a formal proof that a well-known axiomatic characterization of the single-source shortest path problem is correct. Namely, we prove that in a directed graph with a non-negative cost function on the edges the single-source shortest path function is the only function that satisfies a set of four axioms. In the second section, we give a formal proof of the correctness of an axiomatic characterization of the single-source shortest path problem for directed graphs with general cost functions. The axioms here are more involved because we have to account for potential negative cycles in the graph. The axioms are summarized in three Isabelle locales. notify = [Launchbury] title = The Correctness of Launchbury's Natural Semantics for Lazy Evaluation author = Joachim Breitner date = 2013-01-31 topic = Computer science/Programming languages/Lambda calculi, Computer science/Semantics abstract = In his seminal paper "Natural Semantics for Lazy Evaluation", John Launchbury proves his semantics correct with respect to a denotational semantics, and outlines an adequacy proof. We have formalized both semantics and machine-checked the correctness proof, clarifying some details. Furthermore, we provide a new and more direct adequacy proof that does not require intermediate operational semantics. extra-history = Change history: [2014-05-24]: Added the proof of adequacy, as well as simplified and improved the existing proofs. Adjusted abstract accordingly. [2015-03-16]: Booleans and if-then-else added to syntax and semantics, making this entry suitable to be used by the entry "Call_Arity". notify = [Call_Arity] title = The Safety of Call Arity author = Joachim Breitner date = 2015-02-20 topic = Computer science/Programming languages/Transformations abstract = We formalize the Call Arity analysis, as implemented in GHC, and prove both functional correctness and, more interestingly, safety (i.e. the transformation does not increase allocation).

We use syntax and the denotational semantics from the entry "Launchbury", where we formalized Launchbury's natural semantics for lazy evaluation.

The functional correctness of Call Arity is proved with regard to that denotational semantics. The operational properties are shown with regard to a small-step semantics akin to Sestoft's mark 1 machine, which we prove to be equivalent to Launchbury's semantics.

We use Christian Urban's Nominal2 package to define our terms and make use of Brian Huffman's HOLCF package for the domain-theoretical aspects of the development. extra-history = Change history: [2015-03-16]: This entry now builds on top of the Launchbury entry, and the equivalency proof of the natural and the small-step semantics was added. notify = [CCS] title = CCS in nominal logic author = Jesper Bengtson date = 2012-05-29 topic = Computer science/Concurrency/Process calculi abstract = We formalise a large portion of CCS as described in Milner's book 'Communication and Concurrency' using the nominal datatype package in Isabelle. Our results include many of the standard theorems of bisimulation equivalence and congruence, for both weak and strong versions. One main goal of this formalisation is to keep the machine-checked proofs as close to their pen-and-paper counterpart as possible.

This entry is described in detail in Bengtson's thesis. notify = [Pi_Calculus] title = The pi-calculus in nominal logic author = Jesper Bengtson date = 2012-05-29 topic = Computer science/Concurrency/Process calculi abstract = We formalise the pi-calculus using the nominal datatype package, based on ideas from the nominal logic by Pitts et al., and demonstrate an implementation in Isabelle/HOL. The purpose is to derive powerful induction rules for the semantics in order to conduct machine checkable proofs, closely following the intuitive arguments found in manual proofs. In this way we have covered many of the standard theorems of bisimulation equivalence and congruence, both late and early, and both strong and weak in a uniform manner. We thus provide one of the most extensive formalisations of a the pi-calculus ever done inside a theorem prover.

A significant gain in our formulation is that agents are identified up to alpha-equivalence, thereby greatly reducing the arguments about bound names. This is a normal strategy for manual proofs about the pi-calculus, but that kind of hand waving has previously been difficult to incorporate smoothly in an interactive theorem prover. We show how the nominal logic formalism and its support in Isabelle accomplishes this and thus significantly reduces the tedium of conducting completely formal proofs. This improves on previous work using weak higher order abstract syntax since we do not need extra assumptions to filter out exotic terms and can keep all arguments within a familiar first-order logic.

This entry is described in detail in Bengtson's thesis. notify = [Psi_Calculi] title = Psi-calculi in Isabelle author = Jesper Bengtson date = 2012-05-29 topic = Computer science/Concurrency/Process calculi abstract = Psi-calculi are extensions of the pi-calculus, accommodating arbitrary nominal datatypes to represent not only data but also communication channels, assertions and conditions, giving it an expressive power beyond the applied pi-calculus and the concurrent constraint pi-calculus.

We have formalised psi-calculi in the interactive theorem prover Isabelle using its nominal datatype package. One distinctive feature is that the framework needs to treat binding sequences, as opposed to single binders, in an efficient way. While different methods for formalising single binder calculi have been proposed over the last decades, representations for such binding sequences are not very well explored.

The main effort in the formalisation is to keep the machine checked proofs as close to their pen-and-paper counterparts as possible. This includes treating all binding sequences as atomic elements, and creating custom induction and inversion rules that to remove the bulk of manual alpha-conversions.

This entry is described in detail in Bengtson's thesis. notify = [Encodability_Process_Calculi] title = Analysing and Comparing Encodability Criteria for Process Calculi author = Kirstin Peters , Rob van Glabbeek date = 2015-08-10 topic = Computer science/Concurrency/Process calculi abstract = Encodings or the proof of their absence are the main way to compare process calculi. To analyse the quality of encodings and to rule out trivial or meaningless encodings, they are augmented with quality criteria. There exists a bunch of different criteria and different variants of criteria in order to reason in different settings. This leads to incomparable results. Moreover it is not always clear whether the criteria used to obtain a result in a particular setting do indeed fit to this setting. We show how to formally reason about and compare encodability criteria by mapping them on requirements on a relation between source and target terms that is induced by the encoding function. In particular we analyse the common criteria full abstraction, operational correspondence, divergence reflection, success sensitiveness, and respect of barbs; e.g. we analyse the exact nature of the simulation relation (coupled simulation versus bisimulation) that is induced by different variants of operational correspondence. This way we reduce the problem of analysing or comparing encodability criteria to the better understood problem of comparing relations on processes. notify = kirstin.peters@tu-berlin.de [Circus] title = Isabelle/Circus author = Abderrahmane Feliachi , Burkhart Wolff , Marie-Claude Gaudel contributors = Makarius Wenzel date = 2012-05-27 topic = Computer science/Concurrency/Process calculi, Computer science/System description languages abstract = The Circus specification language combines elements for complex data and behavior specifications, using an integration of Z and CSP with a refinement calculus. Its semantics is based on Hoare and He's Unifying Theories of Programming (UTP). Isabelle/Circus is a formalization of the UTP and the Circus language in Isabelle/HOL. It contains proof rules and tactic support that allows for proofs of refinement for Circus processes (involving both data and behavioral aspects).

The Isabelle/Circus environment supports a syntax for the semantic definitions which is close to textbook presentations of Circus. This article contains an extended version of corresponding VSTTE Paper together with the complete formal development of its underlying commented theories. extra-history = Change history: [2014-06-05]: More polishing, shorter proofs, added Circus syntax, added Makarius Wenzel as contributor. notify = [Dijkstra_Shortest_Path] title = Dijkstra's Shortest Path Algorithm author = Benedikt Nordhoff , Peter Lammich topic = Computer science/Algorithms/Graph date = 2012-01-30 abstract = We implement and prove correct Dijkstra's algorithm for the single source shortest path problem, conceived in 1956 by E. Dijkstra. The algorithm is implemented using the data refinement framework for monadic, nondeterministic programs. An efficient implementation is derived using data structures from the Isabelle Collection Framework. notify = lammich@in.tum.de [Refine_Monadic] title = Refinement for Monadic Programs author = Peter Lammich topic = Computer science/Programming languages/Logics date = 2012-01-30 abstract = We provide a framework for program and data refinement in Isabelle/HOL. The framework is based on a nondeterminism-monad with assertions, i.e., the monad carries a set of results or an assertion failure. Recursion is expressed by fixed points. For convenience, we also provide while and foreach combinators.

The framework provides tools to automatize canonical tasks, such as verification condition generation, finding appropriate data refinement relations, and refine an executable program to a form that is accepted by the Isabelle/HOL code generator.

This submission comes with a collection of examples and a user-guide, illustrating the usage of the framework. extra-history = Change history: [2012-04-23] Introduced ordered FOREACH loops
[2012-06] New features: REC_rule_arb and RECT_rule_arb allow for generalizing over variables. prepare_code_thms - command extracts code equations for recursion combinators.
[2012-07] New example: Nested DFS for emptiness check of Buchi-automata with witness.
New feature: fo_rule method to apply resolution using first-order matching. Useful for arg_conf, fun_cong.
[2012-08] Adaptation to ICF v2.
[2012-10-05] Adaptations to include support for Automatic Refinement Framework.
[2013-09] This entry now depends on Automatic Refinement
[2014-06] New feature: vc_solve method to solve verification conditions. Maintenace changes: VCG-rules for nfoldli, improved setup for FOREACH-loops.
[2014-07] Now defining recursion via flat domain. Dropped many single-valued prerequisites. Changed notion of data refinement. In single-valued case, this matches the old notion. In non-single valued case, the new notion allows for more convenient rules. In particular, the new definitions allow for projecting away ghost variables as a refinement step.
[2014-11] New features: le-or-fail relation (leof), modular reasoning about loop invariants. notify = lammich@in.tum.de [Refine_Imperative_HOL] title = The Imperative Refinement Framework author = Peter Lammich notify = lammich@in.tum.de date = 2016-08-08 topic = Computer science/Programming languages/Transformations,Computer science/Data structures abstract = We present the Imperative Refinement Framework (IRF), a tool that supports a stepwise refinement based approach to imperative programs. This entry is based on the material we presented in [ITP-2015, CPP-2016]. It uses the Monadic Refinement Framework as a frontend for the specification of the abstract programs, and Imperative/HOL as a backend to generate executable imperative programs. The IRF comes with tool support to synthesize imperative programs from more abstract, functional ones, using efficient imperative implementations for the abstract data structures. This entry also includes the Imperative Isabelle Collection Framework (IICF), which provides a library of re-usable imperative collection data structures. Moreover, this entry contains a quickstart guide and a reference manual, which provide an introduction to using the IRF for Isabelle/HOL experts. It also provids a collection of (partly commented) practical examples, some highlights being Dijkstra's Algorithm, Nested-DFS, and a generic worklist algorithm with subsumption. Finally, this entry contains benchmark scripts that compare the runtime of some examples against reference implementations of the algorithms in Java and C++. [ITP-2015] Peter Lammich: Refinement to Imperative/HOL. ITP 2015: 253--269 [CPP-2016] Peter Lammich: Refinement based verification of imperative data structures. CPP 2016: 27--36 [Automatic_Refinement] title = Automatic Data Refinement author = Peter Lammich topic = Computer science/Programming languages/Logics date = 2013-10-02 abstract = We present the Autoref tool for Isabelle/HOL, which automatically refines algorithms specified over abstract concepts like maps and sets to algorithms over concrete implementations like red-black-trees, and produces a refinement theorem. It is based on ideas borrowed from relational parametricity due to Reynolds and Wadler. The tool allows for rapid prototyping of verified, executable algorithms. Moreover, it can be configured to fine-tune the result to the user~s needs. Our tool is able to automatically instantiate generic algorithms, which greatly simplifies the implementation of executable data structures.

This AFP-entry provides the basic tool, which is then used by the Refinement and Collection Framework to provide automatic data refinement for the nondeterminism monad and various collection datastructures. notify = lammich@in.tum.de [EdmondsKarp_Maxflow] title = Formalizing the Edmonds-Karp Algorithm author = Peter Lammich , S. Reza Sefidgar<> notify = lammich@in.tum.de date = 2016-08-12 topic = Computer science/Algorithms/Graph abstract = We present a formalization of the Ford-Fulkerson method for computing the maximum flow in a network. Our formal proof closely follows a standard textbook proof, and is accessible even without being an expert in Isabelle/HOL--- the interactive theorem prover used for the formalization. We then use stepwise refinement to obtain the Edmonds-Karp algorithm, and formally prove a bound on its complexity. Further refinement yields a verified implementation, whose execution time compares well to an unverified reference implementation in Java. This entry is based on our ITP-2016 paper with the same title. [VerifyThis2018] title = VerifyThis 2018 - Polished Isabelle Solutions author = Peter Lammich , Simon Wimmer topic = Computer science/Algorithms date = 2018-04-27 notify = lammich@in.tum.de abstract = VerifyThis 2018 was a program verification competition associated with ETAPS 2018. It was the 7th event in the VerifyThis competition series. In this entry, we present polished and completed versions of our solutions that we created during the competition. [PseudoHoops] title = Pseudo Hoops author = George Georgescu <>, Laurentiu Leustean <>, Viorel Preoteasa topic = Mathematics/Algebra date = 2011-09-22 abstract = Pseudo-hoops are algebraic structures introduced by B. Bosbach under the name of complementary semigroups. In this formalization we prove some properties of pseudo-hoops and we define the basic concepts of filter and normal filter. The lattice of normal filters is isomorphic with the lattice of congruences of a pseudo-hoop. We also study some important classes of pseudo-hoops. Bounded Wajsberg pseudo-hoops are equivalent to pseudo-Wajsberg algebras and bounded basic pseudo-hoops are equivalent to pseudo-BL algebras. Some examples of pseudo-hoops are given in the last section of the formalization. notify = viorel.preoteasa@aalto.fi [MonoBoolTranAlgebra] title = Algebra of Monotonic Boolean Transformers author = Viorel Preoteasa topic = Computer science/Programming languages/Logics date = 2011-09-22 abstract = Algebras of imperative programming languages have been successful in reasoning about programs. In general an algebra of programs is an algebraic structure with programs as elements and with program compositions (sequential composition, choice, skip) as algebra operations. Various versions of these algebras were introduced to model partial correctness, total correctness, refinement, demonic choice, and other aspects. We formalize here an algebra which can be used to model total correctness, refinement, demonic and angelic choice. The basic model of this algebra are monotonic Boolean transformers (monotonic functions from a Boolean algebra to itself). notify = viorel.preoteasa@aalto.fi [LatticeProperties] title = Lattice Properties author = Viorel Preoteasa topic = Mathematics/Order date = 2011-09-22 abstract = This formalization introduces and collects some algebraic structures based on lattices and complete lattices for use in other developments. The structures introduced are modular, and lattice ordered groups. In addition to the results proved for the new lattices, this formalization also introduces theorems about latices and complete lattices in general. extra-history = Change history: [2012-01-05]: Removed the theory about distributive complete lattices which is in the standard library now. Added a theory about well founded and transitive relations and a result about fixpoints in complete lattices and well founded relations. Moved the results about conjunctive and disjunctive functions to a new theory. Removed the syntactic classes for inf and sup which are in the standard library now. notify = viorel.preoteasa@aalto.fi [Impossible_Geometry] title = Proving the Impossibility of Trisecting an Angle and Doubling the Cube author = Ralph Romanos , Lawrence C. Paulson topic = Mathematics/Algebra, Mathematics/Geometry date = 2012-08-05 abstract = Squaring the circle, doubling the cube and trisecting an angle, using a compass and straightedge alone, are classic unsolved problems first posed by the ancient Greeks. All three problems were proved to be impossible in the 19th century. The following document presents the proof of the impossibility of solving the latter two problems using Isabelle/HOL, following a proof by Carrega. The proof uses elementary methods: no Galois theory or field extensions. The set of points constructible using a compass and straightedge is defined inductively. Radical expressions, which involve only square roots and arithmetic of rational numbers, are defined, and we find that all constructive points have radical coordinates. Finally, doubling the cube and trisecting certain angles requires solving certain cubic equations that can be proved to have no rational roots. The Isabelle proofs require a great many detailed calculations. notify = ralph.romanos@student.ecp.fr, lp15@cam.ac.uk [IP_Addresses] title = IP Addresses author = Cornelius Diekmann , Julius Michaelis , Lars Hupel notify = diekmann@net.in.tum.de date = 2016-06-28 topic = Computer science/Networks abstract = This entry contains a definition of IP addresses and a library to work with them. Generic IP addresses are modeled as machine words of arbitrary length. Derived from this generic definition, IPv4 addresses are 32bit machine words, IPv6 addresses are 128bit words. Additionally, IPv4 addresses can be represented in dot-decimal notation and IPv6 addresses in (compressed) colon-separated notation. We support toString functions and parsers for both notations. Sets of IP addresses can be represented with a netmask (e.g. 192.168.0.0/255.255.0.0) or in CIDR notation (e.g. 192.168.0.0/16). To provide executable code for set operations on IP address ranges, the library includes a datatype to work on arbitrary intervals of machine words. [Simple_Firewall] title = Simple Firewall author = Cornelius Diekmann , Julius Michaelis , Maximilian Haslbeck notify = diekmann@net.in.tum.de, max.haslbeck@gmx.de date = 2016-08-24 topic = Computer science/Networks abstract = We present a simple model of a firewall. The firewall can accept or drop a packet and can match on interfaces, IP addresses, protocol, and ports. It was designed to feature nice mathematical properties: The type of match expressions was carefully crafted such that the conjunction of two match expressions is only one match expression. This model is too simplistic to mirror all aspects of the real world. In the upcoming entry "Iptables Semantics", we will translate the Linux firewall iptables to this model. For a fixed service (e.g. ssh, http), we provide an algorithm to compute an overview of the firewall's filtering behavior. The algorithm computes minimal service matrices, i.e. graphs which partition the complete IPv4 and IPv6 address space and visualize the allowed accesses between partitions. For a detailed description, see Verified iptables Firewall Analysis, IFIP Networking 2016. [Iptables_Semantics] title = Iptables Semantics author = Cornelius Diekmann , Lars Hupel notify = diekmann@net.in.tum.de, hupel@in.tum.de date = 2016-09-09 topic = Computer science/Networks abstract = We present a big step semantics of the filtering behavior of the Linux/netfilter iptables firewall. We provide algorithms to simplify complex iptables rulests to a simple firewall model (c.f. AFP entry Simple_Firewall) and to verify spoofing protection of a ruleset. Internally, we embed our semantics into ternary logic, ultimately supporting every iptables match condition by abstracting over unknowns. Using this AFP entry and all entries it depends on, we created an easy-to-use, stand-alone haskell tool called fffuu. The tool does not require any input —except for the iptables-save dump of the analyzed firewall— and presents interesting results about the user's ruleset. Real-Word firewall errors have been uncovered, and the correctness of rulesets has been proved, with the help of our tool. [Routing] title = Routing author = Julius Michaelis , Cornelius Diekmann notify = afp@liftm.de date = 2016-08-31 topic = Computer science/Networks abstract = This entry contains definitions for routing with routing tables/longest prefix matching. A routing table entry is modelled as a record of a prefix match, a metric, an output port, and an optional next hop. A routing table is a list of entries, sorted by prefix length and metric. Additionally, a parser and serializer for the output of the ip-route command, a function to create a relation from output port to corresponding destination IP space, and a model of a Linux-style router are included. [KBPs] title = Knowledge-based programs author = Peter Gammie topic = Computer science/Automata and formal languages date = 2011-05-17 abstract = Knowledge-based programs (KBPs) are a formalism for directly relating agents' knowledge and behaviour. Here we present a general scheme for compiling KBPs to executable automata with a proof of correctness in Isabelle/HOL. We develop the algorithm top-down, using Isabelle's locale mechanism to structure these proofs, and show that two classic examples can be synthesised using Isabelle's code generator. extra-history = Change history: [2012-03-06]: Add some more views and revive the code generation. notify = kleing@cse.unsw.edu.au [Tarskis_Geometry] title = The independence of Tarski's Euclidean axiom author = T. J. M. Makarios topic = Mathematics/Geometry date = 2012-10-30 abstract = Tarski's axioms of plane geometry are formalized and, using the standard real Cartesian model, shown to be consistent. A substantial theory of the projective plane is developed. Building on this theory, the Klein-Beltrami model of the hyperbolic plane is defined and shown to satisfy all of Tarski's axioms except his Euclidean axiom; thus Tarski's Euclidean axiom is shown to be independent of his other axioms of plane geometry.

An earlier version of this work was the subject of the author's MSc thesis, which contains natural-language explanations of some of the more interesting proofs. notify = tjm1983@gmail.com [IsaGeoCoq] title = Tarski's Parallel Postulate implies the 5th Postulate of Euclid, the Postulate of Playfair and the original Parallel Postulate of Euclid author = Roland Coghetto topic = Mathematics/Geometry license = LGPL date = 2021-01-31 notify = roland_coghetto@hotmail.com abstract =

The GeoCoq library contains a formalization of geometry using the Coq proof assistant. It contains both proofs about the foundations of geometry and high-level proofs in the same style as in high school. We port a part of the GeoCoq 2.4.0 library to Isabelle/HOL: more precisely, the files Chap02.v to Chap13_3.v, suma.v as well as the associated definitions and some useful files for the demonstration of certain parallel postulates. The synthetic approach of the demonstrations is directly inspired by those contained in GeoCoq. The names of the lemmas and theorems used are kept as far as possible as well as the definitions.

It should be noted that T.J.M. Makarios has done some proofs in Tarski's Geometry. It uses a definition that does not quite coincide with the definition used in Geocoq and here. Furthermore, corresponding definitions in the Poincaré Disc Model development are not identical to those defined in GeoCoq.

In the last part, it is formalized that, in the neutral/absolute space, the axiom of the parallels of Tarski's system implies the Playfair axiom, the 5th postulate of Euclid and Euclid's original parallel postulate. These proofs, which are not constructive, are directly inspired by Pierre Boutry, Charly Gries, Julien Narboux and Pascal Schreck.

[General-Triangle] title = The General Triangle Is Unique author = Joachim Breitner topic = Mathematics/Geometry date = 2011-04-01 abstract = Some acute-angled triangles are special, e.g. right-angled or isoscele triangles. Some are not of this kind, but, without measuring angles, look as if they were. In that sense, there is exactly one general triangle. This well-known fact is proven here formally. notify = mail@joachim-breitner.de [LightweightJava] title = Lightweight Java author = Rok Strniša , Matthew Parkinson topic = Computer science/Programming languages/Language definitions date = 2011-02-07 abstract = A fully-formalized and extensible minimal imperative fragment of Java. notify = rok@strnisa.com [Lower_Semicontinuous] title = Lower Semicontinuous Functions author = Bogdan Grechuk topic = Mathematics/Analysis date = 2011-01-08 abstract = We define the notions of lower and upper semicontinuity for functions from a metric space to the extended real line. We prove that a function is both lower and upper semicontinuous if and only if it is continuous. We also give several equivalent characterizations of lower semicontinuity. In particular, we prove that a function is lower semicontinuous if and only if its epigraph is a closed set. Also, we introduce the notion of the lower semicontinuous hull of an arbitrary function and prove its basic properties. notify = hoelzl@in.tum.de [RIPEMD-160-SPARK] title = RIPEMD-160 author = Fabian Immler topic = Computer science/Programming languages/Static analysis date = 2011-01-10 abstract = This work presents a verification of an implementation in SPARK/ADA of the cryptographic hash-function RIPEMD-160. A functional specification of RIPEMD-160 is given in Isabelle/HOL. Proofs for the verification conditions generated by the static-analysis toolset of SPARK certify the functional correctness of the implementation. extra-history = Change history: [2015-11-09]: Entry is now obsolete, moved to Isabelle distribution. notify = immler@in.tum.de [Regular-Sets] title = Regular Sets and Expressions author = Alexander Krauss , Tobias Nipkow contributors = Manuel Eberl topic = Computer science/Automata and formal languages date = 2010-05-12 abstract = This is a library of constructions on regular expressions and languages. It provides the operations of concatenation, Kleene star and derivative on languages. Regular expressions and their meaning are defined. An executable equivalence checker for regular expressions is verified; it does not need automata but works directly on regular expressions. By mapping regular expressions to binary relations, an automatic and complete proof method for (in)equalities of binary relations over union, concatenation and (reflexive) transitive closure is obtained.

Extended regular expressions with complement and intersection are also defined and an equivalence checker is provided. extra-history = Change history: [2011-08-26]: Christian Urban added a theory about derivatives and partial derivatives of regular expressions
[2012-05-10]: Tobias Nipkow added extended regular expressions
[2012-05-10]: Tobias Nipkow added equivalence checking with partial derivatives notify = nipkow@in.tum.de, krauss@in.tum.de, christian.urban@kcl.ac.uk [Regex_Equivalence] title = Unified Decision Procedures for Regular Expression Equivalence author = Tobias Nipkow , Dmitriy Traytel topic = Computer science/Automata and formal languages date = 2014-01-30 abstract = We formalize a unified framework for verified decision procedures for regular expression equivalence. Five recently published formalizations of such decision procedures (three based on derivatives, two on marked regular expressions) can be obtained as instances of the framework. We discover that the two approaches based on marked regular expressions, which were previously thought to be the same, are different, and one seems to produce uniformly smaller automata. The common framework makes it possible to compare the performance of the different decision procedures in a meaningful way. The formalization is described in a paper of the same name presented at Interactive Theorem Proving 2014. notify = nipkow@in.tum.de, traytel@in.tum.de [MSO_Regex_Equivalence] title = Decision Procedures for MSO on Words Based on Derivatives of Regular Expressions author = Dmitriy Traytel , Tobias Nipkow topic = Computer science/Automata and formal languages, Logic/General logic/Decidability of theories date = 2014-06-12 abstract = Monadic second-order logic on finite words (MSO) is a decidable yet expressive logic into which many decision problems can be encoded. Since MSO formulas correspond to regular languages, equivalence of MSO formulas can be reduced to the equivalence of some regular structures (e.g. automata). We verify an executable decision procedure for MSO formulas that is not based on automata but on regular expressions.

Decision procedures for regular expression equivalence have been formalized before, usually based on Brzozowski derivatives. Yet, for a straightforward embedding of MSO formulas into regular expressions an extension of regular expressions with a projection operation is required. We prove total correctness and completeness of an equivalence checker for regular expressions extended in that way. We also define a language-preserving translation of formulas into regular expressions with respect to two different semantics of MSO.

The formalization is described in this ICFP 2013 functional pearl. notify = traytel@in.tum.de, nipkow@in.tum.de [Formula_Derivatives] title = Derivatives of Logical Formulas author = Dmitriy Traytel topic = Computer science/Automata and formal languages, Logic/General logic/Decidability of theories date = 2015-05-28 abstract = We formalize new decision procedures for WS1S, M2L(Str), and Presburger Arithmetics. Formulas of these logics denote regular languages. Unlike traditional decision procedures, we do not translate formulas into automata (nor into regular expressions), at least not explicitly. Instead we devise notions of derivatives (inspired by Brzozowski derivatives for regular expressions) that operate on formulas directly and compute a syntactic bisimulation using these derivatives. The treatment of Boolean connectives and quantifiers is uniform for all mentioned logics and is abstracted into a locale. This locale is then instantiated by different atomic formulas and their derivatives (which may differ even for the same logic under different encodings of interpretations as formal words).

The WS1S instance is described in the draft paper A Coalgebraic Decision Procedure for WS1S by the author. notify = traytel@in.tum.de [Myhill-Nerode] title = The Myhill-Nerode Theorem Based on Regular Expressions author = Chunhan Wu <>, Xingyuan Zhang <>, Christian Urban contributors = Manuel Eberl topic = Computer science/Automata and formal languages date = 2011-08-26 abstract = There are many proofs of the Myhill-Nerode theorem using automata. In this library we give a proof entirely based on regular expressions, since regularity of languages can be conveniently defined using regular expressions (it is more painful in HOL to define regularity in terms of automata). We prove the first direction of the Myhill-Nerode theorem by solving equational systems that involve regular expressions. For the second direction we give two proofs: one using tagging-functions and another using partial derivatives. We also establish various closure properties of regular languages. Most details of the theories are described in our ITP 2011 paper. notify = christian.urban@kcl.ac.uk [Universal_Turing_Machine] title = Universal Turing Machine author = Jian Xu<>, Xingyuan Zhang<>, Christian Urban , Sebastiaan J. C. Joosten topic = Logic/Computability, Computer science/Automata and formal languages date = 2019-02-08 notify = sjcjoosten@gmail.com, christian.urban@kcl.ac.uk abstract = We formalise results from computability theory: recursive functions, undecidability of the halting problem, and the existence of a universal Turing machine. This formalisation is the AFP entry corresponding to the paper Mechanising Turing Machines and Computability Theory in Isabelle/HOL, ITP 2013. [CYK] title = A formalisation of the Cocke-Younger-Kasami algorithm author = Maksym Bortin date = 2016-04-27 topic = Computer science/Algorithms, Computer science/Automata and formal languages abstract = The theory provides a formalisation of the Cocke-Younger-Kasami algorithm (CYK for short), an approach to solving the word problem for context-free languages. CYK decides if a word is in the languages generated by a context-free grammar in Chomsky normal form. The formalized algorithm is executable. notify = maksym.bortin@nicta.com.au [Boolean_Expression_Checkers] title = Boolean Expression Checkers author = Tobias Nipkow date = 2014-06-08 topic = Computer science/Algorithms, Logic/General logic/Mechanization of proofs abstract = This entry provides executable checkers for the following properties of boolean expressions: satisfiability, tautology and equivalence. Internally, the checkers operate on binary decision trees and are reasonably efficient (for purely functional algorithms). extra-history = Change history: [2015-09-23]: Salomon Sickert added an interface that does not require the usage of the Boolean formula datatype. Furthermore the general Mapping type is used instead of an association list. notify = nipkow@in.tum.de [Presburger-Automata] title = Formalizing the Logic-Automaton Connection author = Stefan Berghofer , Markus Reiter <> date = 2009-12-03 topic = Computer science/Automata and formal languages, Logic/General logic/Decidability of theories abstract = This work presents a formalization of a library for automata on bit strings. It forms the basis of a reflection-based decision procedure for Presburger arithmetic, which is efficiently executable thanks to Isabelle's code generator. With this work, we therefore provide a mechanized proof of a well-known connection between logic and automata theory. The formalization is also described in a publication [TPHOLs 2009]. notify = berghofe@in.tum.de [Functional-Automata] title = Functional Automata author = Tobias Nipkow date = 2004-03-30 topic = Computer science/Automata and formal languages abstract = This theory defines deterministic and nondeterministic automata in a functional representation: the transition function/relation and the finality predicate are just functions. Hence the state space may be infinite. It is shown how to convert regular expressions into such automata. A scanner (generator) is implemented with the help of functional automata: the scanner chops the input up into longest recognized substrings. Finally we also show how to convert a certain subclass of functional automata (essentially the finite deterministic ones) into regular sets. notify = nipkow@in.tum.de [Statecharts] title = Formalizing Statecharts using Hierarchical Automata author = Steffen Helke , Florian Kammüller topic = Computer science/Automata and formal languages date = 2010-08-08 abstract = We formalize in Isabelle/HOL the abtract syntax and a synchronous step semantics for the specification language Statecharts. The formalization is based on Hierarchical Automata which allow a structural decomposition of Statecharts into Sequential Automata. To support the composition of Statecharts, we introduce calculating operators to construct a Hierarchical Automaton in a stepwise manner. Furthermore, we present a complete semantics of Statecharts including a theory of data spaces, which enables the modelling of racing effects. We also adapt CTL for Statecharts to build a bridge for future combinations with model checking. However the main motivation of this work is to provide a sound and complete basis for reasoning on Statecharts. As a central meta theorem we prove that the well-formedness of a Statechart is preserved by the semantics. notify = nipkow@in.tum.de [Stuttering_Equivalence] title = Stuttering Equivalence author = Stephan Merz topic = Computer science/Automata and formal languages date = 2012-05-07 abstract =

Two omega-sequences are stuttering equivalent if they differ only by finite repetitions of elements. Stuttering equivalence is a fundamental concept in the theory of concurrent and distributed systems. Notably, Lamport argues that refinement notions for such systems should be insensitive to finite stuttering. Peled and Wilke showed that all PLTL (propositional linear-time temporal logic) properties that are insensitive to stuttering equivalence can be expressed without the next-time operator. Stuttering equivalence is also important for certain verification techniques such as partial-order reduction for model checking.

We formalize stuttering equivalence in Isabelle/HOL. Our development relies on the notion of stuttering sampling functions that may skip blocks of identical sequence elements. We also encode PLTL and prove the theorem due to Peled and Wilke.

extra-history = Change history: [2013-01-31]: Added encoding of PLTL and proved Peled and Wilke's theorem. Adjusted abstract accordingly. notify = Stephan.Merz@loria.fr [Coinductive_Languages] title = A Codatatype of Formal Languages author = Dmitriy Traytel topic = Computer science/Automata and formal languages date = 2013-11-15 abstract =

We define formal languages as a codataype of infinite trees branching over the alphabet. Each node in such a tree indicates whether the path to this node constitutes a word inside or outside of the language. This codatatype is isormorphic to the set of lists representation of languages, but caters for definitions by corecursion and proofs by coinduction.

Regular operations on languages are then defined by primitive corecursion. A difficulty arises here, since the standard definitions of concatenation and iteration from the coalgebraic literature are not primitively corecursive-they require guardedness up-to union/concatenation. Without support for up-to corecursion, these operation must be defined as a composition of primitive ones (and proved being equal to the standard definitions). As an exercise in coinduction we also prove the axioms of Kleene algebra for the defined regular operations.

Furthermore, a language for context-free grammars given by productions in Greibach normal form and an initial nonterminal is constructed by primitive corecursion, yielding an executable decision procedure for the word problem without further ado.

notify = traytel@in.tum.de [Tree-Automata] title = Tree Automata author = Peter Lammich date = 2009-11-25 topic = Computer science/Automata and formal languages abstract = This work presents a machine-checked tree automata library for Standard-ML, OCaml and Haskell. The algorithms are efficient by using appropriate data structures like RB-trees. The available algorithms for non-deterministic automata include membership query, reduction, intersection, union, and emptiness check with computation of a witness for non-emptiness. The executable algorithms are derived from less-concrete, non-executable algorithms using data-refinement techniques. The concrete data structures are from the Isabelle Collections Framework. Moreover, this work contains a formalization of the class of tree-regular languages and its closure properties under set operations. notify = peter.lammich@uni-muenster.de, nipkow@in.tum.de [Depth-First-Search] title = Depth First Search author = Toshiaki Nishihara <>, Yasuhiko Minamide <> date = 2004-06-24 topic = Computer science/Algorithms/Graph abstract = Depth-first search of a graph is formalized with recdef. It is shown that it visits all of the reachable nodes from a given list of nodes. Executable ML code of depth-first search is obtained using the code generation feature of Isabelle/HOL. notify = lp15@cam.ac.uk, krauss@in.tum.de [FFT] title = Fast Fourier Transform author = Clemens Ballarin date = 2005-10-12 topic = Computer science/Algorithms/Mathematical abstract = We formalise a functional implementation of the FFT algorithm over the complex numbers, and its inverse. Both are shown equivalent to the usual definitions of these operations through Vandermonde matrices. They are also shown to be inverse to each other, more precisely, that composition of the inverse and the transformation yield the identity up to a scalar. notify = ballarin@in.tum.de [Gauss-Jordan-Elim-Fun] title = Gauss-Jordan Elimination for Matrices Represented as Functions author = Tobias Nipkow date = 2011-08-19 topic = Computer science/Algorithms/Mathematical, Mathematics/Algebra abstract = This theory provides a compact formulation of Gauss-Jordan elimination for matrices represented as functions. Its distinctive feature is succinctness. It is not meant for large computations. notify = nipkow@in.tum.de [UpDown_Scheme] title = Verification of the UpDown Scheme author = Johannes Hölzl date = 2015-01-28 topic = Computer science/Algorithms/Mathematical abstract = The UpDown scheme is a recursive scheme used to compute the stiffness matrix on a special form of sparse grids. Usually, when discretizing a Euclidean space of dimension d we need O(n^d) points, for n points along each dimension. Sparse grids are a hierarchical representation where the number of points is reduced to O(n * log(n)^d). One disadvantage of such sparse grids is that the algorithm now operate recursively in the dimensions and levels of the sparse grid.

The UpDown scheme allows us to compute the stiffness matrix on such a sparse grid. The stiffness matrix represents the influence of each representation function on the L^2 scalar product. For a detailed description see Dirk Pflüger's PhD thesis. This formalization was developed as an interdisciplinary project (IDP) at the Technische Universität München. notify = hoelzl@in.tum.de [GraphMarkingIBP] title = Verification of the Deutsch-Schorr-Waite Graph Marking Algorithm using Data Refinement author = Viorel Preoteasa , Ralph-Johan Back date = 2010-05-28 topic = Computer science/Algorithms/Graph abstract = The verification of the Deutsch-Schorr-Waite graph marking algorithm is used as a benchmark in many formalizations of pointer programs. The main purpose of this mechanization is to show how data refinement of invariant based programs can be used in verifying practical algorithms. The verification starts with an abstract algorithm working on a graph given by a relation next on nodes. Gradually the abstract program is refined into Deutsch-Schorr-Waite graph marking algorithm where only one bit per graph node of additional memory is used for marking. extra-history = Change history: [2012-01-05]: Updated for the new definition of data refinement and the new syntax for demonic and angelic update statements notify = viorel.preoteasa@aalto.fi [Efficient-Mergesort] title = Efficient Mergesort topic = Computer science/Algorithms date = 2011-11-09 author = Christian Sternagel abstract = We provide a formalization of the mergesort algorithm as used in GHC's Data.List module, proving correctness and stability. Furthermore, experimental data suggests that generated (Haskell-)code for this algorithm is much faster than for previous algorithms available in the Isabelle distribution. extra-history = Change history: [2012-10-24]: Added reference to journal article.
[2018-09-17]: Added theory Efficient_Mergesort that works exclusively with the mutual induction schemas generated by the function package.
[2018-09-19]: Added theory Mergesort_Complexity that proves an upper bound on the number of comparisons that are required by mergesort.
[2018-09-19]: Theory Efficient_Mergesort replaces theory Efficient_Sort but keeping the old name Efficient_Sort. [2020-11-20]: Additional theory Natural_Mergesort that developes an efficient mergesort algorithm without key-functions for educational purposes. notify = c.sternagel@gmail.com [SATSolverVerification] title = Formal Verification of Modern SAT Solvers author = Filip Marić date = 2008-07-23 topic = Computer science/Algorithms abstract = This document contains formal correctness proofs of modern SAT solvers. Following (Krstic et al, 2007) and (Nieuwenhuis et al., 2006), solvers are described using state-transition systems. Several different SAT solver descriptions are given and their partial correctness and termination is proved. These include:

  • a solver based on classical DPLL procedure (using only a backtrack-search with unit propagation),
  • a very general solver with backjumping and learning (similar to the description given in (Nieuwenhuis et al., 2006)), and
  • a solver with a specific conflict analysis algorithm (similar to the description given in (Krstic et al., 2007)).
Within the SAT solver correctness proofs, a large number of lemmas about propositional logic and CNF formulae are proved. This theory is self-contained and could be used for further exploring of properties of CNF based SAT algorithms. notify = [Transitive-Closure] title = Executable Transitive Closures of Finite Relations topic = Computer science/Algorithms/Graph date = 2011-03-14 author = Christian Sternagel , René Thiemann license = LGPL abstract = We provide a generic work-list algorithm to compute the transitive closure of finite relations where only successors of newly detected states are generated. This algorithm is then instantiated for lists over arbitrary carriers and red black trees (which are faster but require a linear order on the carrier), respectively. Our formalization was performed as part of the IsaFoR/CeTA project where reflexive transitive closures of large tree automata have to be computed. extra-history = Change history: [2014-09-04] added example simprocs in Finite_Transitive_Closure_Simprocs notify = c.sternagel@gmail.com, rene.thiemann@uibk.ac.at [Transitive-Closure-II] title = Executable Transitive Closures topic = Computer science/Algorithms/Graph date = 2012-02-29 author = René Thiemann license = LGPL abstract =

We provide a generic work-list algorithm to compute the (reflexive-)transitive closure of relations where only successors of newly detected states are generated. In contrast to our previous work, the relations do not have to be finite, but each element must only have finitely many (indirect) successors. Moreover, a subsumption relation can be used instead of pure equality. An executable variant of the algorithm is available where the generic operations are instantiated with list operations.

This formalization was performed as part of the IsaFoR/CeTA project, and it has been used to certify size-change termination proofs where large transitive closures have to be computed.

notify = rene.thiemann@uibk.ac.at [MuchAdoAboutTwo] title = Much Ado About Two author = Sascha Böhme date = 2007-11-06 topic = Computer science/Algorithms abstract = This article is an Isabelle formalisation of a paper with the same title. In a similar way as Knuth's 0-1-principle for sorting algorithms, that paper develops a 0-1-2-principle for parallel prefix computations. notify = boehmes@in.tum.de [DiskPaxos] title = Proving the Correctness of Disk Paxos date = 2005-06-22 author = Mauro Jaskelioff , Stephan Merz topic = Computer science/Algorithms/Distributed abstract = Disk Paxos is an algorithm for building arbitrary fault-tolerant distributed systems. The specification of Disk Paxos has been proved correct informally and tested using the TLC model checker, but up to now, it has never been fully formally verified. In this work we have formally verified its correctness using the Isabelle theorem prover and the HOL logic system, showing that Isabelle is a practical tool for verifying properties of TLA+ specifications. notify = kleing@cse.unsw.edu.au [GenClock] title = Formalization of a Generalized Protocol for Clock Synchronization author = Alwen Tiu date = 2005-06-24 topic = Computer science/Algorithms/Distributed abstract = We formalize the generalized Byzantine fault-tolerant clock synchronization protocol of Schneider. This protocol abstracts from particular algorithms or implementations for clock synchronization. This abstraction includes several assumptions on the behaviors of physical clocks and on general properties of concrete algorithms/implementations. Based on these assumptions the correctness of the protocol is proved by Schneider. His proof was later verified by Shankar using the theorem prover EHDM (precursor to PVS). Our formalization in Isabelle/HOL is based on Shankar's formalization. notify = kleing@cse.unsw.edu.au [ClockSynchInst] title = Instances of Schneider's generalized protocol of clock synchronization author = Damián Barsotti date = 2006-03-15 topic = Computer science/Algorithms/Distributed abstract = F. B. Schneider ("Understanding protocols for Byzantine clock synchronization") generalizes a number of protocols for Byzantine fault-tolerant clock synchronization and presents a uniform proof for their correctness. In Schneider's schema, each processor maintains a local clock by periodically adjusting each value to one computed by a convergence function applied to the readings of all the clocks. Then, correctness of an algorithm, i.e. that the readings of two clocks at any time are within a fixed bound of each other, is based upon some conditions on the convergence function. To prove that a particular clock synchronization algorithm is correct it suffices to show that the convergence function used by the algorithm meets Schneider's conditions. Using the theorem prover Isabelle, we formalize the proofs that the convergence functions of two algorithms, namely, the Interactive Convergence Algorithm (ICA) of Lamport and Melliar-Smith and the Fault-tolerant Midpoint algorithm of Lundelius-Lynch, meet Schneider's conditions. Furthermore, we experiment on handling some parts of the proofs with fully automatic tools like ICS and CVC-lite. These theories are part of a joint work with Alwen Tiu and Leonor P. Nieto "Verification of Clock Synchronization Algorithms: Experiments on a combination of deductive tools" in proceedings of AVOCS 2005. In this work the correctness of Schneider schema was also verified using Isabelle (entry GenClock in AFP). notify = kleing@cse.unsw.edu.au [Heard_Of] title = Verifying Fault-Tolerant Distributed Algorithms in the Heard-Of Model date = 2012-07-27 author = Henri Debrat , Stephan Merz topic = Computer science/Algorithms/Distributed abstract = Distributed computing is inherently based on replication, promising increased tolerance to failures of individual computing nodes or communication channels. Realizing this promise, however, involves quite subtle algorithmic mechanisms, and requires precise statements about the kinds and numbers of faults that an algorithm tolerates (such as process crashes, communication faults or corrupted values). The landmark theorem due to Fischer, Lynch, and Paterson shows that it is impossible to achieve Consensus among N asynchronously communicating nodes in the presence of even a single permanent failure. Existing solutions must rely on assumptions of "partial synchrony".

Indeed, there have been numerous misunderstandings on what exactly a given algorithm is supposed to realize in what kinds of environments. Moreover, the abundance of subtly different computational models complicates comparisons between different algorithms. Charron-Bost and Schiper introduced the Heard-Of model for representing algorithms and failure assumptions in a uniform framework, simplifying comparisons between algorithms.

In this contribution, we represent the Heard-Of model in Isabelle/HOL. We define two semantics of runs of algorithms with different unit of atomicity and relate these through a reduction theorem that allows us to verify algorithms in the coarse-grained semantics (where proofs are easier) and infer their correctness for the fine-grained one (which corresponds to actual executions). We instantiate the framework by verifying six Consensus algorithms that differ in the underlying algorithmic mechanisms and the kinds of faults they tolerate. notify = Stephan.Merz@loria.fr [Consensus_Refined] title = Consensus Refined date = 2015-03-18 author = Ognjen Maric <>, Christoph Sprenger topic = Computer science/Algorithms/Distributed abstract = Algorithms for solving the consensus problem are fundamental to distributed computing. Despite their brevity, their ability to operate in concurrent, asynchronous and failure-prone environments comes at the cost of complex and subtle behaviors. Accordingly, understanding how they work and proving their correctness is a non-trivial endeavor where abstraction is immensely helpful. Moreover, research on consensus has yielded a large number of algorithms, many of which appear to share common algorithmic ideas. A natural question is whether and how these similarities can be distilled and described in a precise, unified way. In this work, we combine stepwise refinement and lockstep models to provide an abstract and unified view of a sizeable family of consensus algorithms. Our models provide insights into the design choices underlying the different algorithms, and classify them based on those choices. notify = sprenger@inf.ethz.ch [Key_Agreement_Strong_Adversaries] title = Refining Authenticated Key Agreement with Strong Adversaries author = Joseph Lallemand , Christoph Sprenger topic = Computer science/Security license = LGPL date = 2017-01-31 notify = joseph.lallemand@loria.fr, sprenger@inf.ethz.ch abstract = We develop a family of key agreement protocols that are correct by construction. Our work substantially extends prior work on developing security protocols by refinement. First, we strengthen the adversary by allowing him to compromise different resources of protocol participants, such as their long-term keys or their session keys. This enables the systematic development of protocols that ensure strong properties such as perfect forward secrecy. Second, we broaden the class of protocols supported to include those with non-atomic keys and equationally defined cryptographic operators. We use these extensions to develop key agreement protocols including signed Diffie-Hellman and the core of IKEv1 and SKEME. [Security_Protocol_Refinement] title = Developing Security Protocols by Refinement author = Christoph Sprenger , Ivano Somaini<> topic = Computer science/Security license = LGPL date = 2017-05-24 notify = sprenger@inf.ethz.ch abstract = We propose a development method for security protocols based on stepwise refinement. Our refinement strategy transforms abstract security goals into protocols that are secure when operating over an insecure channel controlled by a Dolev-Yao-style intruder. As intermediate levels of abstraction, we employ messageless guard protocols and channel protocols communicating over channels with security properties. These abstractions provide insights on why protocols are secure and foster the development of families of protocols sharing common structure and properties. We have implemented our method in Isabelle/HOL and used it to develop different entity authentication and key establishment protocols, including realistic features such as key confirmation, replay caches, and encrypted tickets. Our development highlights that guard protocols and channel protocols provide fundamental abstractions for bridging the gap between security properties and standard protocol descriptions based on cryptographic messages. It also shows that our refinement approach scales to protocols of nontrivial size and complexity. [Abortable_Linearizable_Modules] title = Abortable Linearizable Modules author = Rachid Guerraoui , Viktor Kuncak , Giuliano Losa date = 2012-03-01 topic = Computer science/Algorithms/Distributed abstract = We define the Abortable Linearizable Module automaton (ALM for short) and prove its key composition property using the IOA theory of HOLCF. The ALM is at the heart of the Speculative Linearizability framework. This framework simplifies devising correct speculative algorithms by enabling their decomposition into independent modules that can be analyzed and proved correct in isolation. It is particularly useful when working in a distributed environment, where the need to tolerate faults and asynchrony has made current monolithic protocols so intricate that it is no longer tractable to check their correctness. Our theory contains a typical example of a refinement proof in the I/O-automata framework of Lynch and Tuttle. notify = giuliano@losa.fr, nipkow@in.tum.de [Amortized_Complexity] title = Amortized Complexity Verified author = Tobias Nipkow date = 2014-07-07 topic = Computer science/Data structures abstract = A framework for the analysis of the amortized complexity of functional data structures is formalized in Isabelle/HOL and applied to a number of standard examples and to the folowing non-trivial ones: skew heaps, splay trees, splay heaps and pairing heaps.

A preliminary version of this work (without pairing heaps) is described in a paper published in the proceedings of the conference on Interactive Theorem Proving ITP 2015. An extended version of this publication is available here. extra-history = Change history: [2015-03-17]: Added pairing heaps by Hauke Brinkop.
[2016-07-12]: Moved splay heaps from here to Splay_Tree
[2016-07-14]: Moved pairing heaps from here to the new Pairing_Heap notify = nipkow@in.tum.de [Dynamic_Tables] title = Parameterized Dynamic Tables author = Tobias Nipkow date = 2015-06-07 topic = Computer science/Data structures abstract = This article formalizes the amortized analysis of dynamic tables parameterized with their minimal and maximal load factors and the expansion and contraction factors.

A full description is found in a companion paper. notify = nipkow@in.tum.de [AVL-Trees] title = AVL Trees author = Tobias Nipkow , Cornelia Pusch <> date = 2004-03-19 topic = Computer science/Data structures abstract = Two formalizations of AVL trees with room for extensions. The first formalization is monolithic and shorter, the second one in two stages, longer and a bit simpler. The final implementation is the same. If you are interested in developing this further, please contact gerwin.klein@nicta.com.au. extra-history = Change history: [2011-04-11]: Ondrej Kuncar added delete function notify = kleing@cse.unsw.edu.au [BDD] title = BDD Normalisation author = Veronika Ortner <>, Norbert Schirmer <> date = 2008-02-29 topic = Computer science/Data structures abstract = We present the verification of the normalisation of a binary decision diagram (BDD). The normalisation follows the original algorithm presented by Bryant in 1986 and transforms an ordered BDD in a reduced, ordered and shared BDD. The verification is based on Hoare logics. notify = kleing@cse.unsw.edu.au, norbert.schirmer@web.de [BinarySearchTree] title = Binary Search Trees author = Viktor Kuncak date = 2004-04-05 topic = Computer science/Data structures abstract = The correctness is shown of binary search tree operations (lookup, insert and remove) implementing a set. Two versions are given, for both structured and linear (tactic-style) proofs. An implementation of integer-indexed maps is also verified. notify = lp15@cam.ac.uk [Splay_Tree] title = Splay Tree author = Tobias Nipkow notify = nipkow@in.tum.de date = 2014-08-12 topic = Computer science/Data structures abstract = Splay trees are self-adjusting binary search trees which were invented by Sleator and Tarjan [JACM 1985]. This entry provides executable and verified functional splay trees as well as the related splay heaps (due to Okasaki).

The amortized complexity of splay trees and heaps is analyzed in the AFP entry Amortized Complexity. extra-history = Change history: [2016-07-12]: Moved splay heaps here from Amortized_Complexity [Root_Balanced_Tree] title = Root-Balanced Tree author = Tobias Nipkow notify = nipkow@in.tum.de date = 2017-08-20 topic = Computer science/Data structures abstract =

Andersson introduced general balanced trees, search trees based on the design principle of partial rebuilding: perform update operations naively until the tree becomes too unbalanced, at which point a whole subtree is rebalanced. This article defines and analyzes a functional version of general balanced trees, which we call root-balanced trees. Using a lightweight model of execution time, amortized logarithmic complexity is verified in the theorem prover Isabelle.

This is the Isabelle formalization of the material decribed in the APLAS 2017 article Verified Root-Balanced Trees by the same author, which also presents experimental results that show competitiveness of root-balanced with AVL and red-black trees.

[Skew_Heap] title = Skew Heap author = Tobias Nipkow date = 2014-08-13 topic = Computer science/Data structures abstract = Skew heaps are an amazingly simple and lightweight implementation of priority queues. They were invented by Sleator and Tarjan [SIAM 1986] and have logarithmic amortized complexity. This entry provides executable and verified functional skew heaps.

The amortized complexity of skew heaps is analyzed in the AFP entry Amortized Complexity. notify = nipkow@in.tum.de [Pairing_Heap] title = Pairing Heap author = Hauke Brinkop , Tobias Nipkow date = 2016-07-14 topic = Computer science/Data structures abstract = This library defines three different versions of pairing heaps: a functional version of the original design based on binary trees [Fredman et al. 1986], the version by Okasaki [1998] and a modified version of the latter that is free of structural invariants.

The amortized complexity of pairing heaps is analyzed in the AFP article Amortized Complexity. extra-0 = Origin: This library was extracted from Amortized Complexity and extended. notify = nipkow@in.tum.de [Priority_Queue_Braun] title = Priority Queues Based on Braun Trees author = Tobias Nipkow date = 2014-09-04 topic = Computer science/Data structures abstract = This entry verifies priority queues based on Braun trees. Insertion and deletion take logarithmic time and preserve the balanced nature of Braun trees. Two implementations of deletion are provided. notify = nipkow@in.tum.de extra-history = Change history: [2019-12-16]: Added theory Priority_Queue_Braun2 with second version of del_min [Binomial-Queues] title = Functional Binomial Queues author = René Neumann date = 2010-10-28 topic = Computer science/Data structures abstract = Priority queues are an important data structure and efficient implementations of them are crucial. We implement a functional variant of binomial queues in Isabelle/HOL and show its functional correctness. A verification against an abstract reference specification of priority queues has also been attempted, but could not be achieved to the full extent. notify = florian.haftmann@informatik.tu-muenchen.de [Binomial-Heaps] title = Binomial Heaps and Skew Binomial Heaps author = Rene Meis , Finn Nielsen , Peter Lammich date = 2010-10-28 topic = Computer science/Data structures abstract = We implement and prove correct binomial heaps and skew binomial heaps. Both are data-structures for priority queues. While binomial heaps have logarithmic findMin, deleteMin, insert, and meld operations, skew binomial heaps have constant time findMin, insert, and meld operations, and only the deleteMin-operation is logarithmic. This is achieved by using skew links to avoid cascading linking on insert-operations, and data-structural bootstrapping to get constant-time findMin and meld operations. Our implementation follows the paper by Brodal and Okasaki. notify = peter.lammich@uni-muenster.de [Finger-Trees] title = Finger Trees author = Benedikt Nordhoff , Stefan Körner , Peter Lammich date = 2010-10-28 topic = Computer science/Data structures abstract = We implement and prove correct 2-3 finger trees. Finger trees are a general purpose data structure, that can be used to efficiently implement other data structures, such as priority queues. Intuitively, a finger tree is an annotated sequence, where the annotations are elements of a monoid. Apart from operations to access the ends of the sequence, the main operation is to split the sequence at the point where a monotone predicate over the sum of the left part of the sequence becomes true for the first time. The implementation follows the paper of Hinze and Paterson. The code generator can be used to get efficient, verified code. notify = peter.lammich@uni-muenster.de [Trie] title = Trie author = Andreas Lochbihler , Tobias Nipkow date = 2015-03-30 topic = Computer science/Data structures abstract = This article formalizes the ``trie'' data structure invented by Fredkin [CACM 1960]. It also provides a specialization where the entries in the trie are lists. extra-0 = Origin: This article was extracted from existing articles by the authors. notify = nipkow@in.tum.de [FinFun] title = Code Generation for Functions as Data author = Andreas Lochbihler date = 2009-05-06 topic = Computer science/Data structures abstract = FinFuns are total functions that are constant except for a finite set of points, i.e. a generalisation of finite maps. They are formalised as a new type in Isabelle/HOL such that the code generator can handle equality tests and quantification on FinFuns. On the code output level, FinFuns are explicitly represented by constant functions and pointwise updates, similarly to associative lists. Inside the logic, they behave like ordinary functions with extensionality. Via the update/constant pattern, a recursion combinator and an induction rule for FinFuns allow for defining and reasoning about operators on FinFun that are also executable. extra-history = Change history: [2010-08-13]: new concept domain of a FinFun as a FinFun (revision 34b3517cbc09)
[2010-11-04]: new conversion function from FinFun to list of elements in the domain (revision 0c167102e6ed)
[2012-03-07]: replace sets as FinFuns by predicates as FinFuns because the set type constructor has been reintroduced (revision b7aa87989f3a) notify = nipkow@in.tum.de [Collections] title = Collections Framework author = Peter Lammich contributors = Andreas Lochbihler , Thomas Tuerk <> date = 2009-11-25 topic = Computer science/Data structures abstract = This development provides an efficient, extensible, machine checked collections framework. The library adopts the concepts of interface, implementation and generic algorithm from object-oriented programming and implements them in Isabelle/HOL. The framework features the use of data refinement techniques to refine an abstract specification (using high-level concepts like sets) to a more concrete implementation (using collection datastructures, like red-black-trees). The code-generator of Isabelle/HOL can be used to generate efficient code. extra-history = Change history: [2010-10-08]: New Interfaces: OrderedSet, OrderedMap, List. Fifo now implements list-interface: Function names changed: put/get --> enqueue/dequeue. New Implementations: ArrayList, ArrayHashMap, ArrayHashSet, TrieMap, TrieSet. Invariant-free datastructures: Invariant implicitely hidden in typedef. Record-interfaces: All operations of an interface encapsulated as record. Examples moved to examples subdirectory.
[2010-12-01]: New Interfaces: Priority Queues, Annotated Lists. Implemented by finger trees, (skew) binomial queues.
[2011-10-10]: SetSpec: Added operations: sng, isSng, bexists, size_abort, diff, filter, iterate_rule_insertP MapSpec: Added operations: sng, isSng, iterate_rule_insertP, bexists, size, size_abort, restrict, map_image_filter, map_value_image_filter Some maintenance changes
[2012-04-25]: New iterator foundation by Tuerk. Various maintenance changes.
[2012-08]: Collections V2. New features: Polymorphic iterators. Generic algorithm instantiation where required. Naming scheme changed from xx_opname to xx.opname. A compatibility file CollectionsV1 tries to simplify porting of existing theories, by providing old naming scheme and the old monomorphic iterator locales.
[2013-09]: Added Generic Collection Framework based on Autoref. The GenCF provides: Arbitrary nesting, full integration with Autoref.
[2014-06]: Maintenace changes to GenCF: Optimized inj_image on list_set. op_set_cart (Cartesian product). big-Union operation. atLeastLessThan - operation ({a..<b})
notify = lammich@in.tum.de [Containers] title = Light-weight Containers author = Andreas Lochbihler contributors = René Thiemann date = 2013-04-15 topic = Computer science/Data structures abstract = This development provides a framework for container types like sets and maps such that generated code implements these containers with different (efficient) data structures. Thanks to type classes and refinement during code generation, this light-weight approach can seamlessly replace Isabelle's default setup for code generation. Heuristics automatically pick one of the available data structures depending on the type of elements to be stored, but users can also choose on their own. The extensible design permits to add more implementations at any time.

To support arbitrary nesting of sets, we define a linear order on sets based on a linear order of the elements and provide efficient implementations. It even allows to compare complements with non-complements. extra-history = Change history: [2013-07-11]: add pretty printing for sets (revision 7f3f52c5f5fa)
[2013-09-20]: provide generators for canonical type class instantiations (revision 159f4401f4a8 by René Thiemann)
[2014-07-08]: add support for going from partial functions to mappings (revision 7a6fc957e8ed)
[2018-03-05]: add two application examples: depth-first search and 2SAT (revision e5e1a1da2411) notify = mail@andreas-lochbihler.de [FileRefinement] title = File Refinement author = Karen Zee , Viktor Kuncak date = 2004-12-09 topic = Computer science/Data structures abstract = These theories illustrates the verification of basic file operations (file creation, file read and file write) in the Isabelle theorem prover. We describe a file at two levels of abstraction: an abstract file represented as a resizable array, and a concrete file represented using data blocks. notify = kkz@mit.edu [Datatype_Order_Generator] title = Generating linear orders for datatypes author = René Thiemann date = 2012-08-07 topic = Computer science/Data structures abstract = We provide a framework for registering automatic methods to derive class instances of datatypes, as it is possible using Haskell's ``deriving Ord, Show, ...'' feature.

We further implemented such automatic methods to derive (linear) orders or hash-functions which are required in the Isabelle Collection Framework. Moreover, for the tactic of Huffman and Krauss to show that a datatype is countable, we implemented a wrapper so that this tactic becomes accessible in our framework.

Our formalization was performed as part of the IsaFoR/CeTA project. With our new tactic we could completely remove tedious proofs for linear orders of two datatypes.

This development is aimed at datatypes generated by the "old_datatype" command. notify = rene.thiemann@uibk.ac.at [Deriving] title = Deriving class instances for datatypes author = Christian Sternagel , René Thiemann date = 2015-03-11 topic = Computer science/Data structures abstract =

We provide a framework for registering automatic methods to derive class instances of datatypes, as it is possible using Haskell's ``deriving Ord, Show, ...'' feature.

We further implemented such automatic methods to derive comparators, linear orders, parametrizable equality functions, and hash-functions which are required in the Isabelle Collection Framework and the Container Framework. Moreover, for the tactic of Blanchette to show that a datatype is countable, we implemented a wrapper so that this tactic becomes accessible in our framework. All of the generators are based on the infrastructure that is provided by the BNF-based datatype package.

Our formalization was performed as part of the IsaFoR/CeTA project. With our new tactics we could remove several tedious proofs for (conditional) linear orders, and conditional equality operators within IsaFoR and the Container Framework.

notify = rene.thiemann@uibk.ac.at [List-Index] title = List Index date = 2010-02-20 author = Tobias Nipkow topic = Computer science/Data structures abstract = This theory provides functions for finding the index of an element in a list, by predicate and by value. notify = nipkow@in.tum.de [List-Infinite] title = Infinite Lists date = 2011-02-23 author = David Trachtenherz <> topic = Computer science/Data structures abstract = We introduce a theory of infinite lists in HOL formalized as functions over naturals (folder ListInf, theories ListInf and ListInf_Prefix). It also provides additional results for finite lists (theory ListInf/List2), natural numbers (folder CommonArith, esp. division/modulo, naturals with infinity), sets (folder CommonSet, esp. cutting/truncating sets, traversing sets of naturals). notify = nipkow@in.tum.de [Matrix] title = Executable Matrix Operations on Matrices of Arbitrary Dimensions topic = Computer science/Data structures date = 2010-06-17 author = Christian Sternagel , René Thiemann license = LGPL abstract = We provide the operations of matrix addition, multiplication, transposition, and matrix comparisons as executable functions over ordered semirings. Moreover, it is proven that strongly normalizing (monotone) orders can be lifted to strongly normalizing (monotone) orders over matrices. We further show that the standard semirings over the naturals, integers, and rationals, as well as the arctic semirings satisfy the axioms that are required by our matrix theory. Our formalization is part of the CeTA system which contains several termination techniques. The provided theories have been essential to formalize matrix-interpretations and arctic interpretations. extra-history = Change history: [2010-09-17]: Moved theory on arbitrary (ordered) semirings to Abstract Rewriting. notify = rene.thiemann@uibk.ac.at, christian.sternagel@uibk.ac.at [Matrix_Tensor] title = Tensor Product of Matrices topic = Computer science/Data structures, Mathematics/Algebra date = 2016-01-18 author = T.V.H. Prathamesh abstract = In this work, the Kronecker tensor product of matrices and the proofs of some of its properties are formalized. Properties which have been formalized include associativity of the tensor product and the mixed-product property. notify = prathamesh@imsc.res.in [Huffman] title = The Textbook Proof of Huffman's Algorithm author = Jasmin Christian Blanchette date = 2008-10-15 topic = Computer science/Data structures abstract = Huffman's algorithm is a procedure for constructing a binary tree with minimum weighted path length. This report presents a formal proof of the correctness of Huffman's algorithm written using Isabelle/HOL. Our proof closely follows the sketches found in standard algorithms textbooks, uncovering a few snags in the process. Another distinguishing feature of our formalization is the use of custom induction rules to help Isabelle's automatic tactics, leading to very short proofs for most of the lemmas. notify = jasmin.blanchette@gmail.com [Partial_Function_MR] title = Mutually Recursive Partial Functions author = René Thiemann topic = Computer science/Functional programming date = 2014-02-18 license = LGPL abstract = We provide a wrapper around the partial-function command that supports mutual recursion. notify = rene.thiemann@uibk.ac.at [Lifting_Definition_Option] title = Lifting Definition Option author = René Thiemann topic = Computer science/Functional programming date = 2014-10-13 license = LGPL abstract = We implemented a command that can be used to easily generate elements of a restricted type {x :: 'a. P x}, provided the definition is of the form f ys = (if check ys then Some(generate ys :: 'a) else None) where ys is a list of variables y1 ... yn and check ys ==> P(generate ys) can be proved.

In principle, such a definition is also directly possible using the lift_definition command. However, then this definition will not be suitable for code-generation. To this end, we automated a more complex construction of Joachim Breitner which is amenable for code-generation, and where the test check ys will only be performed once. In the automation, one auxiliary type is created, and Isabelle's lifting- and transfer-package is invoked several times. notify = rene.thiemann@uibk.ac.at [Coinductive] title = Coinductive topic = Computer science/Functional programming author = Andreas Lochbihler contributors = Johannes Hölzl date = 2010-02-12 abstract = This article collects formalisations of general-purpose coinductive data types and sets. Currently, it contains coinductive natural numbers, coinductive lists, i.e. lazy lists or streams, infinite streams, coinductive terminated lists, coinductive resumptions, a library of operations on coinductive lists, and a version of König's lemma as an application for coinductive lists.
The initial theory was contributed by Paulson and Wenzel. Extensions and other coinductive formalisations of general interest are welcome. extra-history = Change history: [2010-06-10]: coinductive lists: setup for quotient package (revision 015574f3bf3c)
[2010-06-28]: new codatatype terminated lazy lists (revision e12de475c558)
[2010-08-04]: terminated lazy lists: setup for quotient package; more lemmas (revision 6ead626f1d01)
[2010-08-17]: Koenig's lemma as an example application for coinductive lists (revision f81ce373fa96)
[2011-02-01]: lazy implementation of coinductive (terminated) lists for the code generator (revision 6034973dce83)
[2011-07-20]: new codatatype resumption (revision 811364c776c7)
[2012-06-27]: new codatatype stream with operations (with contributions by Peter Gammie) (revision dd789a56473c)
[2013-03-13]: construct codatatypes with the BNF package and adjust the definitions and proofs, setup for lifting and transfer packages (revision f593eda5b2c0)
[2013-09-20]: stream theory uses type and operations from HOL/BNF/Examples/Stream (revision 692809b2b262)
[2014-04-03]: ccpo structure on codatatypes used to define ldrop, ldropWhile, lfilter, lconcat as least fixpoint; ccpo topology on coinductive lists contributed by Johannes Hölzl; added examples (revision 23cd8156bd42)
notify = mail@andreas-lochbihler.de [Stream-Fusion] title = Stream Fusion author = Brian Huffman topic = Computer science/Functional programming date = 2009-04-29 abstract = Stream Fusion is a system for removing intermediate list structures from Haskell programs; it consists of a Haskell library along with several compiler rewrite rules. (The library is available online.)

These theories contain a formalization of much of the Stream Fusion library in HOLCF. Lazy list and stream types are defined, along with coercions between the two types, as well as an equivalence relation for streams that generate the same list. List and stream versions of map, filter, foldr, enumFromTo, append, zipWith, and concatMap are defined, and the stream versions are shown to respect stream equivalence. notify = brianh@cs.pdx.edu [Tycon] title = Type Constructor Classes and Monad Transformers author = Brian Huffman date = 2012-06-26 topic = Computer science/Functional programming abstract = These theories contain a formalization of first class type constructors and axiomatic constructor classes for HOLCF. This work is described in detail in the ICFP 2012 paper Formal Verification of Monad Transformers by the author. The formalization is a revised and updated version of earlier joint work with Matthews and White.

Based on the hierarchy of type classes in Haskell, we define classes for functors, monads, monad-plus, etc. Each one includes all the standard laws as axioms. We also provide a new user command, tycondef, for defining new type constructors in HOLCF. Using tycondef, we instantiate the type class hierarchy with various monads and monad transformers. notify = huffman@in.tum.de [CoreC++] title = CoreC++ author = Daniel Wasserrab date = 2006-05-15 topic = Computer science/Programming languages/Language definitions abstract = We present an operational semantics and type safety proof for multiple inheritance in C++. The semantics models the behavior of method calls, field accesses, and two forms of casts in C++ class hierarchies. For explanations see the OOPSLA 2006 paper by Wasserrab, Nipkow, Snelting and Tip. notify = nipkow@in.tum.de [FeatherweightJava] title = A Theory of Featherweight Java in Isabelle/HOL author = J. Nathan Foster , Dimitrios Vytiniotis date = 2006-03-31 topic = Computer science/Programming languages/Language definitions abstract = We formalize the type system, small-step operational semantics, and type soundness proof for Featherweight Java, a simple object calculus, in Isabelle/HOL. notify = kleing@cse.unsw.edu.au [Jinja] title = Jinja is not Java author = Gerwin Klein , Tobias Nipkow date = 2005-06-01 topic = Computer science/Programming languages/Language definitions abstract = We introduce Jinja, a Java-like programming language with a formal semantics designed to exhibit core features of the Java language architecture. Jinja is a compromise between realism of the language and tractability and clarity of the formal semantics. The following aspects are formalised: a big and a small step operational semantics for Jinja and a proof of their equivalence; a type system and a definite initialisation analysis; a type safety proof of the small step semantics; a virtual machine (JVM), its operational semantics and its type system; a type safety proof for the JVM; a bytecode verifier, i.e. data flow analyser for the JVM; a correctness proof of the bytecode verifier w.r.t. the type system; a compiler and a proof that it preserves semantics and well-typedness. The emphasis of this work is not on particular language features but on providing a unified model of the source language, the virtual machine and the compiler. The whole development has been carried out in the theorem prover Isabelle/HOL. notify = kleing@cse.unsw.edu.au, nipkow@in.tum.de [JinjaThreads] title = Jinja with Threads author = Andreas Lochbihler date = 2007-12-03 topic = Computer science/Programming languages/Language definitions abstract = We extend the Jinja source code semantics by Klein and Nipkow with Java-style arrays and threads. Concurrency is captured in a generic framework semantics for adding concurrency through interleaving to a sequential semantics, which features dynamic thread creation, inter-thread communication via shared memory, lock synchronisation and joins. Also, threads can suspend themselves and be notified by others. We instantiate the framework with the adapted versions of both Jinja source and byte code and show type safety for the multithreaded case. Equally, the compiler from source to byte code is extended, for which we prove weak bisimilarity between the source code small step semantics and the defensive Jinja virtual machine. On top of this, we formalise the JMM and show the DRF guarantee and consistency. For description of the different parts, see Lochbihler's papers at FOOL 2008, ESOP 2010, ITP 2011, and ESOP 2012. extra-history = Change history: [2008-04-23]: added bytecode formalisation with arrays and threads, added thread joins (revision f74a8be156a7)
[2009-04-27]: added verified compiler from source code to bytecode; encapsulate native methods in separate semantics (revision e4f26541e58a)
[2009-11-30]: extended compiler correctness proof to infinite and deadlocking computations (revision e50282397435)
[2010-06-08]: added thread interruption; new abstract memory model with sequential consistency as implementation (revision 0cb9e8dbd78d)
[2010-06-28]: new thread interruption model (revision c0440d0a1177)
[2010-10-15]: preliminary version of the Java memory model for source code (revision 02fee0ef3ca2)
[2010-12-16]: improved version of the Java memory model, also for bytecode executable scheduler for source code semantics (revision 1f41c1842f5a)
[2011-02-02]: simplified code generator setup new random scheduler (revision 3059dafd013f)
[2011-07-21]: new interruption model, generalized JMM proof of DRF guarantee, allow class Object to declare methods and fields, simplified subtyping relation, corrected division and modulo implementation (revision 46e4181ed142)
[2012-02-16]: added example programs (revision bf0b06c8913d)
[2012-11-21]: type safety proof for the Java memory model, allow spurious wake-ups (revision 76063d860ae0)
[2013-05-16]: support for non-deterministic memory allocators (revision cc3344a49ced)
[2017-10-20]: add an atomic compare-and-swap operation for volatile fields (revision a6189b1d6b30)
notify = mail@andreas-lochbihler.de [Locally-Nameless-Sigma] title = Locally Nameless Sigma Calculus author = Ludovic Henrio , Florian Kammüller , Bianca Lutz , Henry Sudhof date = 2010-04-30 topic = Computer science/Programming languages/Language definitions abstract = We present a Theory of Objects based on the original functional sigma-calculus by Abadi and Cardelli but with an additional parameter to methods. We prove confluence of the operational semantics following the outline of Nipkow's proof of confluence for the lambda-calculus reusing his theory Commutation, a generic diamond lemma reduction. We furthermore formalize a simple type system for our sigma-calculus including a proof of type safety. The entire development uses the concept of Locally Nameless representation for binders. We reuse an earlier proof of confluence for a simpler sigma-calculus based on de Bruijn indices and lists to represent objects. notify = nipkow@in.tum.de [Attack_Trees] title = Attack Trees in Isabelle for GDPR compliance of IoT healthcare systems author = Florian Kammueller topic = Computer science/Security date = 2020-04-27 notify = florian.kammuller@gmail.com abstract = In this article, we present a proof theory for Attack Trees. Attack Trees are a well established and useful model for the construction of attacks on systems since they allow a stepwise exploration of high level attacks in application scenarios. Using the expressiveness of Higher Order Logic in Isabelle, we develop a generic theory of Attack Trees with a state-based semantics based on Kripke structures and CTL. The resulting framework allows mechanically supported logic analysis of the meta-theory of the proof calculus of Attack Trees and at the same time the developed proof theory enables application to case studies. A central correctness and completeness result proved in Isabelle establishes a connection between the notion of Attack Tree validity and CTL. The application is illustrated on the example of a healthcare IoT system and GDPR compliance verification. [AutoFocus-Stream] title = AutoFocus Stream Processing for Single-Clocking and Multi-Clocking Semantics author = David Trachtenherz <> date = 2011-02-23 topic = Computer science/Programming languages/Language definitions abstract = We formalize the AutoFocus Semantics (a time-synchronous subset of the Focus formalism) as stream processing functions on finite and infinite message streams represented as finite/infinite lists. The formalization comprises both the conventional single-clocking semantics (uniform global clock for all components and communications channels) and its extension to multi-clocking semantics (internal execution clocking of a component may be a multiple of the external communication clocking). The semantics is defined by generic stream processing functions making it suitable for simulation/code generation in Isabelle/HOL. Furthermore, a number of AutoFocus semantics properties are formalized using definitions from the IntervalLogic theories. notify = nipkow@in.tum.de [FocusStreamsCaseStudies] title = Stream Processing Components: Isabelle/HOL Formalisation and Case Studies author = Maria Spichkova date = 2013-11-14 topic = Computer science/Programming languages/Language definitions abstract = This set of theories presents an Isabelle/HOL formalisation of stream processing components introduced in Focus, a framework for formal specification and development of interactive systems. This is an extended and updated version of the formalisation, which was elaborated within the methodology "Focus on Isabelle". In addition, we also applied the formalisation on three case studies that cover different application areas: process control (Steam Boiler System), data transmission (FlexRay communication protocol), memory and processing components (Automotive-Gateway System). notify = lp15@cam.ac.uk, maria.spichkova@rmit.edu.au [Isabelle_Meta_Model] title = A Meta-Model for the Isabelle API author = Frédéric Tuong , Burkhart Wolff date = 2015-09-16 topic = Computer science/Programming languages/Language definitions abstract = We represent a theory of (a fragment of) Isabelle/HOL in Isabelle/HOL. The purpose of this exercise is to write packages for domain-specific specifications such as class models, B-machines, ..., and generally speaking, any domain-specific languages whose abstract syntax can be defined by a HOL "datatype". On this basis, the Isabelle code-generator can then be used to generate code for global context transformations as well as tactic code.

Consequently the package is geared towards parsing, printing and code-generation to the Isabelle API. It is at the moment not sufficiently rich for doing meta theory on Isabelle itself. Extensions in this direction are possible though.

Moreover, the chosen fragment is fairly rudimentary. However it should be easily adapted to one's needs if a package is written on top of it. The supported API contains types, terms, transformation of global context like definitions and data-type declarations as well as infrastructure for Isar-setups.

This theory is drawn from the Featherweight OCL project where it is used to construct a package for object-oriented data-type theories generated from UML class diagrams. The Featherweight OCL, for example, allows for both the direct execution of compiled tactic code by the Isabelle API as well as the generation of ".thy"-files for debugging purposes.

Gained experience from this project shows that the compiled code is sufficiently efficient for practical purposes while being based on a formal model on which properties of the package can be proven such as termination of certain transformations, correctness, etc. notify = tuong@users.gforge.inria.fr, wolff@lri.fr [Clean] title = Clean - An Abstract Imperative Programming Language and its Theory author = Frédéric Tuong , Burkhart Wolff topic = Computer science/Programming languages, Computer science/Semantics date = 2019-10-04 notify = wolff@lri.fr, ftuong@lri.fr abstract = Clean is based on a simple, abstract execution model for an imperative target language. “Abstract” is understood in contrast to “Concrete Semantics”; alternatively, the term “shallow-style embedding” could be used. It strives for a type-safe notion of program-variables, an incremental construction of the typed state-space, support of incremental verification, and open-world extensibility of new type definitions being intertwined with the program definitions. Clean is based on a “no-frills” state-exception monad with the usual definitions of bind and unit for the compositional glue of state-based computations. Clean offers conditionals and loops supporting C-like control-flow operators such as break and return. The state-space construction is based on the extensible record package. Direct recursion of procedures is supported. Clean’s design strives for extreme simplicity. It is geared towards symbolic execution and proven correct verification tools. The underlying libraries of this package, however, deliberately restrict themselves to the most elementary infrastructure for these tasks. The package is intended to serve as demonstrator semantic backend for Isabelle/C, or for the test-generation techniques. [PCF] title = Logical Relations for PCF author = Peter Gammie date = 2012-07-01 topic = Computer science/Programming languages/Lambda calculi abstract = We apply Andy Pitts's methods of defining relations over domains to several classical results in the literature. We show that the Y combinator coincides with the domain-theoretic fixpoint operator, that parallel-or and the Plotkin existential are not definable in PCF, that the continuation semantics for PCF coincides with the direct semantics, and that our domain-theoretic semantics for PCF is adequate for reasoning about contextual equivalence in an operational semantics. Our version of PCF is untyped and has both strict and non-strict function abstractions. The development is carried out in HOLCF. notify = peteg42@gmail.com [POPLmark-deBruijn] title = POPLmark Challenge Via de Bruijn Indices author = Stefan Berghofer date = 2007-08-02 topic = Computer science/Programming languages/Lambda calculi abstract = We present a solution to the POPLmark challenge designed by Aydemir et al., which has as a goal the formalization of the meta-theory of System F<:. The formalization is carried out in the theorem prover Isabelle/HOL using an encoding based on de Bruijn indices. We start with a relatively simple formalization covering only the basic features of System F<:, and explain how it can be extended to also cover records and more advanced binding constructs. notify = berghofe@in.tum.de [Lam-ml-Normalization] title = Strong Normalization of Moggis's Computational Metalanguage author = Christian Doczkal date = 2010-08-29 topic = Computer science/Programming languages/Lambda calculi abstract = Handling variable binding is one of the main difficulties in formal proofs. In this context, Moggi's computational metalanguage serves as an interesting case study. It features monadic types and a commuting conversion rule that rearranges the binding structure. Lindley and Stark have given an elegant proof of strong normalization for this calculus. The key construction in their proof is a notion of relational TT-lifting, using stacks of elimination contexts to obtain a Girard-Tait style logical relation. I give a formalization of their proof in Isabelle/HOL-Nominal with a particular emphasis on the treatment of bound variables. notify = doczkal@ps.uni-saarland.de, nipkow@in.tum.de [MiniML] title = Mini ML author = Wolfgang Naraschewski <>, Tobias Nipkow date = 2004-03-19 topic = Computer science/Programming languages/Type systems abstract = This theory defines the type inference rules and the type inference algorithm W for MiniML (simply-typed lambda terms with let) due to Milner. It proves the soundness and completeness of W w.r.t. the rules. notify = kleing@cse.unsw.edu.au [Simpl] title = A Sequential Imperative Programming Language Syntax, Semantics, Hoare Logics and Verification Environment author = Norbert Schirmer <> date = 2008-02-29 topic = Computer science/Programming languages/Language definitions, Computer science/Programming languages/Logics license = LGPL abstract = We present the theory of Simpl, a sequential imperative programming language. We introduce its syntax, its semantics (big and small-step operational semantics) and Hoare logics for both partial as well as total correctness. We prove soundness and completeness of the Hoare logic. We integrate and automate the Hoare logic in Isabelle/HOL to obtain a practically usable verification environment for imperative programs. Simpl is independent of a concrete programming language but expressive enough to cover all common language features: mutually recursive procedures, abrupt termination and exceptions, runtime faults, local and global variables, pointers and heap, expressions with side effects, pointers to procedures, partial application and closures, dynamic method invocation and also unbounded nondeterminism. notify = kleing@cse.unsw.edu.au, norbert.schirmer@web.de [Separation_Algebra] title = Separation Algebra author = Gerwin Klein , Rafal Kolanski , Andrew Boyton date = 2012-05-11 topic = Computer science/Programming languages/Logics license = BSD abstract = We present a generic type class implementation of separation algebra for Isabelle/HOL as well as lemmas and generic tactics which can be used directly for any instantiation of the type class.

The ex directory contains example instantiations that include structures such as a heap or virtual memory.

The abstract separation algebra is based upon "Abstract Separation Logic" by Calcagno et al. These theories are also the basis of the ITP 2012 rough diamond "Mechanised Separation Algebra" by the authors.

The aim of this work is to support and significantly reduce the effort for future separation logic developments in Isabelle/HOL by factoring out the part of separation logic that can be treated abstractly once and for all. This includes developing typical default rule sets for reasoning as well as automated tactic support for separation logic. notify = kleing@cse.unsw.edu.au, rafal.kolanski@nicta.com.au [Separation_Logic_Imperative_HOL] title = A Separation Logic Framework for Imperative HOL author = Peter Lammich , Rene Meis date = 2012-11-14 topic = Computer science/Programming languages/Logics license = BSD abstract = We provide a framework for separation-logic based correctness proofs of Imperative HOL programs. Our framework comes with a set of proof methods to automate canonical tasks such as verification condition generation and frame inference. Moreover, we provide a set of examples that show the applicability of our framework. The examples include algorithms on lists, hash-tables, and union-find trees. We also provide abstract interfaces for lists, maps, and sets, that allow to develop generic imperative algorithms and use data-refinement techniques.
As we target Imperative HOL, our programs can be translated to efficiently executable code in various target languages, including ML, OCaml, Haskell, and Scala. notify = lammich@in.tum.de [Inductive_Confidentiality] title = Inductive Study of Confidentiality author = Giampaolo Bella date = 2012-05-02 topic = Computer science/Security abstract = This document contains the full theory files accompanying article Inductive Study of Confidentiality --- for Everyone in Formal Aspects of Computing. They aim at an illustrative and didactic presentation of the Inductive Method of protocol analysis, focusing on the treatment of one of the main goals of security protocols: confidentiality against a threat model. The treatment of confidentiality, which in fact forms a key aspect of all protocol analysis tools, has been found cryptic by many learners of the Inductive Method, hence the motivation for this work. The theory files in this document guide the reader step by step towards design and proof of significant confidentiality theorems. These are developed against two threat models, the standard Dolev-Yao and a more audacious one, the General Attacker, which turns out to be particularly useful also for teaching purposes. notify = giamp@dmi.unict.it [Possibilistic_Noninterference] title = Possibilistic Noninterference author = Andrei Popescu , Johannes Hölzl date = 2012-09-10 topic = Computer science/Security, Computer science/Programming languages/Type systems abstract = We formalize a wide variety of Volpano/Smith-style noninterference notions for a while language with parallel composition. We systematize and classify these notions according to compositionality w.r.t. the language constructs. Compositionality yields sound syntactic criteria (a.k.a. type systems) in a uniform way.

An article about these proofs is published in the proceedings of the conference Certified Programs and Proofs 2012. notify = hoelzl@in.tum.de [SIFUM_Type_Systems] title = A Formalization of Assumptions and Guarantees for Compositional Noninterference author = Sylvia Grewe , Heiko Mantel , Daniel Schoepe date = 2014-04-23 topic = Computer science/Security, Computer science/Programming languages/Type systems abstract = Research in information-flow security aims at developing methods to identify undesired information leaks within programs from private (high) sources to public (low) sinks. For a concurrent system, it is desirable to have compositional analysis methods that allow for analyzing each thread independently and that nevertheless guarantee that the parallel composition of successfully analyzed threads satisfies a global security guarantee. However, such a compositional analysis should not be overly pessimistic about what an environment might do with shared resources. Otherwise, the analysis will reject many intuitively secure programs.

The paper "Assumptions and Guarantees for Compositional Noninterference" by Mantel et. al. presents one solution for this problem: an approach for compositionally reasoning about non-interference in concurrent programs via rely-guarantee-style reasoning. We present an Isabelle/HOL formalization of the concepts and proofs of this approach. notify = [Dependent_SIFUM_Type_Systems] title = A Dependent Security Type System for Concurrent Imperative Programs author = Toby Murray , Robert Sison<>, Edward Pierzchalski<>, Christine Rizkallah notify = toby.murray@unimelb.edu.au date = 2016-06-25 topic = Computer science/Security, Computer science/Programming languages/Type systems abstract = The paper "Compositional Verification and Refinement of Concurrent Value-Dependent Noninterference" by Murray et. al. (CSF 2016) presents a dependent security type system for compositionally verifying a value-dependent noninterference property, defined in (Murray, PLAS 2015), for concurrent programs. This development formalises that security definition, the type system and its soundness proof, and demonstrates its application on some small examples. It was derived from the SIFUM_Type_Systems AFP entry, by Sylvia Grewe, Heiko Mantel and Daniel Schoepe, and whose structure it inherits. extra-history = Change history: [2016-08-19]: Removed unused "stop" parameter and "stop_no_eval" assumption from the sifum_security locale. (revision dbc482d36372) [2016-09-27]: Added security locale support for the imposition of requirements on the initial memory. (revision cce4ceb74ddb) [Dependent_SIFUM_Refinement] title = Compositional Security-Preserving Refinement for Concurrent Imperative Programs author = Toby Murray , Robert Sison<>, Edward Pierzchalski<>, Christine Rizkallah notify = toby.murray@unimelb.edu.au date = 2016-06-28 topic = Computer science/Security abstract = The paper "Compositional Verification and Refinement of Concurrent Value-Dependent Noninterference" by Murray et. al. (CSF 2016) presents a compositional theory of refinement for a value-dependent noninterference property, defined in (Murray, PLAS 2015), for concurrent programs. This development formalises that refinement theory, and demonstrates its application on some small examples. extra-history = Change history: [2016-08-19]: Removed unused "stop" parameters from the sifum_refinement locale. (revision dbc482d36372) [2016-09-02]: TobyM extended "simple" refinement theory to be usable for all bisimulations. (revision 547f31c25f60) [Relational-Incorrectness-Logic] title = An Under-Approximate Relational Logic author = Toby Murray topic = Computer science/Programming languages/Logics, Computer science/Security date = 2020-03-12 notify = toby.murray@unimelb.edu.au abstract = Recently, authors have proposed under-approximate logics for reasoning about programs. So far, all such logics have been confined to reasoning about individual program behaviours. Yet there exist many over-approximate relational logics for reasoning about pairs of programs and relating their behaviours. We present the first under-approximate relational logic, for the simple imperative language IMP. We prove our logic is both sound and complete. Additionally, we show how reasoning in this logic can be decomposed into non-relational reasoning in an under-approximate Hoare logic, mirroring Beringer’s result for over-approximate relational logics. We illustrate the application of our logic on some small examples in which we provably demonstrate the presence of insecurity. [Strong_Security] title = A Formalization of Strong Security author = Sylvia Grewe , Alexander Lux , Heiko Mantel , Jens Sauer date = 2014-04-23 topic = Computer science/Security, Computer science/Programming languages/Type systems abstract = Research in information-flow security aims at developing methods to identify undesired information leaks within programs from private sources to public sinks. Noninterference captures this intuition. Strong security from Sabelfeld and Sands formalizes noninterference for concurrent systems.

We present an Isabelle/HOL formalization of strong security for arbitrary security lattices (Sabelfeld and Sands use a two-element security lattice in the original publication). The formalization includes compositionality proofs for strong security and a soundness proof for a security type system that checks strong security for programs in a simple while language with dynamic thread creation.

Our formalization of the security type system is abstract in the language for expressions and in the semantic side conditions for expressions. It can easily be instantiated with different syntactic approximations for these side conditions. The soundness proof of such an instantiation boils down to showing that these syntactic approximations imply the semantic side conditions. notify = [WHATandWHERE_Security] title = A Formalization of Declassification with WHAT-and-WHERE-Security author = Sylvia Grewe , Alexander Lux , Heiko Mantel , Jens Sauer date = 2014-04-23 topic = Computer science/Security, Computer science/Programming languages/Type systems abstract = Research in information-flow security aims at developing methods to identify undesired information leaks within programs from private sources to public sinks. Noninterference captures this intuition by requiring that no information whatsoever flows from private sources to public sinks. However, in practice this definition is often too strict: Depending on the intuitive desired security policy, the controlled declassification of certain private information (WHAT) at certain points in the program (WHERE) might not result in an undesired information leak.

We present an Isabelle/HOL formalization of such a security property for controlled declassification, namely WHAT&WHERE-security from "Scheduler-Independent Declassification" by Lux, Mantel, and Perner. The formalization includes compositionality proofs for and a soundness proof for a security type system that checks for programs in a simple while language with dynamic thread creation.

Our formalization of the security type system is abstract in the language for expressions and in the semantic side conditions for expressions. It can easily be instantiated with different syntactic approximations for these side conditions. The soundness proof of such an instantiation boils down to showing that these syntactic approximations imply the semantic side conditions.

This Isabelle/HOL formalization uses theories from the entry Strong Security. notify = [VolpanoSmith] title = A Correctness Proof for the Volpano/Smith Security Typing System author = Gregor Snelting , Daniel Wasserrab date = 2008-09-02 topic = Computer science/Programming languages/Type systems, Computer science/Security abstract = The Volpano/Smith/Irvine security type systems requires that variables are annotated as high (secret) or low (public), and provides typing rules which guarantee that secret values cannot leak to public output ports. This property of a program is called confidentiality. For a simple while-language without threads, our proof shows that typeability in the Volpano/Smith system guarantees noninterference. Noninterference means that if two initial states for program execution are low-equivalent, then the final states are low-equivalent as well. This indeed implies that secret values cannot leak to public ports. The proof defines an abstract syntax and operational semantics for programs, formalizes noninterference, and then proceeds by rule induction on the operational semantics. The mathematically most intricate part is the treatment of implicit flows. Note that the Volpano/Smith system is not flow-sensitive and thus quite unprecise, resulting in false alarms. However, due to the correctness property, all potential breaks of confidentiality are discovered. notify = [Abstract-Hoare-Logics] title = Abstract Hoare Logics author = Tobias Nipkow date = 2006-08-08 topic = Computer science/Programming languages/Logics abstract = These therories describe Hoare logics for a number of imperative language constructs, from while-loops to mutually recursive procedures. Both partial and total correctness are treated. In particular a proof system for total correctness of recursive procedures in the presence of unbounded nondeterminism is presented. notify = nipkow@in.tum.de [Stone_Algebras] title = Stone Algebras author = Walter Guttmann notify = walter.guttmann@canterbury.ac.nz date = 2016-09-06 topic = Mathematics/Order abstract = A range of algebras between lattices and Boolean algebras generalise the notion of a complement. We develop a hierarchy of these pseudo-complemented algebras that includes Stone algebras. Independently of this theory we study filters based on partial orders. Both theories are combined to prove Chen and Grätzer's construction theorem for Stone algebras. The latter involves extensive reasoning about algebraic structures in addition to reasoning in algebraic structures. [Kleene_Algebra] title = Kleene Algebra author = Alasdair Armstrong <>, Georg Struth , Tjark Weber date = 2013-01-15 topic = Computer science/Programming languages/Logics, Computer science/Automata and formal languages, Mathematics/Algebra abstract = These files contain a formalisation of variants of Kleene algebras and their most important models as axiomatic type classes in Isabelle/HOL. Kleene algebras are foundational structures in computing with applications ranging from automata and language theory to computational modeling, program construction and verification.

We start with formalising dioids, which are additively idempotent semirings, and expand them by axiomatisations of the Kleene star for finite iteration and an omega operation for infinite iteration. We show that powersets over a given monoid, (regular) languages, sets of paths in a graph, sets of computation traces, binary relations and formal power series form Kleene algebras, and consider further models based on lattices, max-plus semirings and min-plus semirings. We also demonstrate that dioids are closed under the formation of matrices (proofs for Kleene algebras remain to be completed).

On the one hand we have aimed at a reference formalisation of variants of Kleene algebras that covers a wide range of variants and the core theorems in a structured and modular way and provides readable proofs at text book level. On the other hand, we intend to use this algebraic hierarchy and its models as a generic algebraic middle-layer from which programming applications can quickly be explored, implemented and verified. notify = g.struth@sheffield.ac.uk, tjark.weber@it.uu.se [KAT_and_DRA] title = Kleene Algebra with Tests and Demonic Refinement Algebras author = Alasdair Armstrong <>, Victor B. F. Gomes , Georg Struth date = 2014-01-23 topic = Computer science/Programming languages/Logics, Computer science/Automata and formal languages, Mathematics/Algebra abstract = We formalise Kleene algebra with tests (KAT) and demonic refinement algebra (DRA) in Isabelle/HOL. KAT is relevant for program verification and correctness proofs in the partial correctness setting. While DRA targets similar applications in the context of total correctness. Our formalisation contains the two most important models of these algebras: binary relations in the case of KAT and predicate transformers in the case of DRA. In addition, we derive the inference rules for Hoare logic in KAT and its relational model and present a simple formally verified program verification tool prototype based on the algebraic approach. notify = g.struth@dcs.shef.ac.uk [KAD] title = Kleene Algebras with Domain author = Victor B. F. Gomes , Walter Guttmann , Peter Höfner , Georg Struth , Tjark Weber date = 2016-04-12 topic = Computer science/Programming languages/Logics, Computer science/Automata and formal languages, Mathematics/Algebra abstract = Kleene algebras with domain are Kleene algebras endowed with an operation that maps each element of the algebra to its domain of definition (or its complement) in abstract fashion. They form a simple algebraic basis for Hoare logics, dynamic logics or predicate transformer semantics. We formalise a modular hierarchy of algebras with domain and antidomain (domain complement) operations in Isabelle/HOL that ranges from domain and antidomain semigroups to modal Kleene algebras and divergence Kleene algebras. We link these algebras with models of binary relations and program traces. We include some examples from modal logics, termination and program analysis. notify = walter.guttman@canterbury.ac.nz, g.struth@sheffield.ac.uk, tjark.weber@it.uu.se [Regular_Algebras] title = Regular Algebras author = Simon Foster , Georg Struth date = 2014-05-21 topic = Computer science/Automata and formal languages, Mathematics/Algebra abstract = Regular algebras axiomatise the equational theory of regular expressions as induced by regular language identity. We use Isabelle/HOL for a detailed systematic study of regular algebras given by Boffa, Conway, Kozen and Salomaa. We investigate the relationships between these classes, formalise a soundness proof for the smallest class (Salomaa's) and obtain completeness of the largest one (Boffa's) relative to a deep result by Krob. In addition we provide a large collection of regular identities in the general setting of Boffa's axiom. Our regular algebra hierarchy is orthogonal to the Kleene algebra hierarchy in the Archive of Formal Proofs; we have not aimed at an integration for pragmatic reasons. notify = simon.foster@york.ac.uk, g.struth@sheffield.ac.uk [BytecodeLogicJmlTypes] title = A Bytecode Logic for JML and Types author = Lennart Beringer <>, Martin Hofmann date = 2008-12-12 topic = Computer science/Programming languages/Logics abstract = This document contains the Isabelle/HOL sources underlying the paper A bytecode logic for JML and types by Beringer and Hofmann, updated to Isabelle 2008. We present a program logic for a subset of sequential Java bytecode that is suitable for representing both, features found in high-level specification language JML as well as interpretations of high-level type systems. To this end, we introduce a fine-grained collection of assertions, including strong invariants, local annotations and VDM-reminiscent partial-correctness specifications. Thanks to a goal-oriented structure and interpretation of judgements, verification may proceed without recourse to an additional control flow analysis. The suitability for interpreting intensional type systems is illustrated by the proof-carrying-code style encoding of a type system for a first-order functional language which guarantees a constant upper bound on the number of objects allocated throughout an execution, be the execution terminating or non-terminating. Like the published paper, the formal development is restricted to a comparatively small subset of the JVML, lacking (among other features) exceptions, arrays, virtual methods, and static fields. This shortcoming has been overcome meanwhile, as our paper has formed the basis of the Mobius base logic, a program logic for the full sequential fragment of the JVML. Indeed, the present formalisation formed the basis of a subsequent formalisation of the Mobius base logic in the proof assistant Coq, which includes a proof of soundness with respect to the Bicolano operational semantics by Pichardie. notify = [DataRefinementIBP] title = Semantics and Data Refinement of Invariant Based Programs author = Viorel Preoteasa , Ralph-Johan Back date = 2010-05-28 topic = Computer science/Programming languages/Logics abstract = The invariant based programming is a technique of constructing correct programs by first identifying the basic situations (pre- and post-conditions and invariants) that can occur during the execution of the program, and then defining the transitions and proving that they preserve the invariants. Data refinement is a technique of building correct programs working on concrete datatypes as refinements of more abstract programs. In the theories presented here we formalize the predicate transformer semantics for invariant based programs and their data refinement. extra-history = Change history: [2012-01-05]: Moved some general complete lattice properties to the AFP entry Lattice Properties. Changed the definition of the data refinement relation to be more general and updated all corresponding theorems. Added new syntax for demonic and angelic update statements. notify = viorel.preoteasa@aalto.fi [RefinementReactive] title = Formalization of Refinement Calculus for Reactive Systems author = Viorel Preoteasa date = 2014-10-08 topic = Computer science/Programming languages/Logics abstract = We present a formalization of refinement calculus for reactive systems. Refinement calculus is based on monotonic predicate transformers (monotonic functions from sets of post-states to sets of pre-states), and it is a powerful formalism for reasoning about imperative programs. We model reactive systems as monotonic property transformers that transform sets of output infinite sequences into sets of input infinite sequences. Within this semantics we can model refinement of reactive systems, (unbounded) angelic and demonic nondeterminism, sequential composition, and other semantic properties. We can model systems that may fail for some inputs, and we can model compatibility of systems. We can specify systems that have liveness properties using linear temporal logic, and we can refine system specifications into systems based on symbolic transitions systems, suitable for implementations. notify = viorel.preoteasa@aalto.fi [SIFPL] title = Secure information flow and program logics author = Lennart Beringer <>, Martin Hofmann date = 2008-11-10 topic = Computer science/Programming languages/Logics, Computer science/Security abstract = We present interpretations of type systems for secure information flow in Hoare logic, complementing previous encodings in relational program logics. We first treat the imperative language IMP, extended by a simple procedure call mechanism. For this language we consider base-line non-interference in the style of Volpano et al. and the flow-sensitive type system by Hunt and Sands. In both cases, we show how typing derivations may be used to automatically generate proofs in the program logic that certify the absence of illicit flows. We then add instructions for object creation and manipulation, and derive appropriate proof rules for base-line non-interference. As a consequence of our work, standard verification technology may be used for verifying that a concrete program satisfies the non-interference property.

The present proof development represents an update of the formalisation underlying our paper [CSF 2007] and is intended to resolve any ambiguities that may be present in the paper. notify = lennart.beringer@ifi.lmu.de [TLA] title = A Definitional Encoding of TLA* in Isabelle/HOL author = Gudmund Grov , Stephan Merz date = 2011-11-19 topic = Computer science/Programming languages/Logics abstract = We mechanise the logic TLA* [Merz 1999], an extension of Lamport's Temporal Logic of Actions (TLA) [Lamport 1994] for specifying and reasoning about concurrent and reactive systems. Aiming at a framework for mechanising] the verification of TLA (or TLA*) specifications, this contribution reuses some elements from a previous axiomatic encoding of TLA in Isabelle/HOL by the second author [Merz 1998], which has been part of the Isabelle distribution. In contrast to that previous work, we give here a shallow, definitional embedding, with the following highlights:

  • a theory of infinite sequences, including a formalisation of the concepts of stuttering invariance central to TLA and TLA*;
  • a definition of the semantics of TLA*, which extends TLA by a mutually-recursive definition of formulas and pre-formulas, generalising TLA action formulas;
  • a substantial set of derived proof rules, including the TLA* axioms and Lamport's proof rules for system verification;
  • a set of examples illustrating the usage of Isabelle/TLA* for reasoning about systems.
Note that this work is unrelated to the ongoing development of a proof system for the specification language TLA+, which includes an encoding of TLA+ as a new Isabelle object logic [Chaudhuri et al 2010]. notify = ggrov@inf.ed.ac.uk [Compiling-Exceptions-Correctly] title = Compiling Exceptions Correctly author = Tobias Nipkow date = 2004-07-09 topic = Computer science/Programming languages/Compiling abstract = An exception compilation scheme that dynamically creates and removes exception handler entries on the stack. A formalization of an article of the same name by Hutton and Wright. notify = nipkow@in.tum.de [NormByEval] title = Normalization by Evaluation author = Klaus Aehlig , Tobias Nipkow date = 2008-02-18 topic = Computer science/Programming languages/Compiling abstract = This article formalizes normalization by evaluation as implemented in Isabelle. Lambda calculus plus term rewriting is compiled into a functional program with pattern matching. It is proved that the result of a successful evaluation is a) correct, i.e. equivalent to the input, and b) in normal form. notify = nipkow@in.tum.de [Program-Conflict-Analysis] title = Formalization of Conflict Analysis of Programs with Procedures, Thread Creation, and Monitors topic = Computer science/Programming languages/Static analysis author = Peter Lammich , Markus Müller-Olm date = 2007-12-14 abstract = In this work we formally verify the soundness and precision of a static program analysis that detects conflicts (e. g. data races) in programs with procedures, thread creation and monitors with the Isabelle theorem prover. As common in static program analysis, our program model abstracts guarded branching by nondeterministic branching, but completely interprets the call-/return behavior of procedures, synchronization by monitors, and thread creation. The analysis is based on the observation that all conflicts already occur in a class of particularly restricted schedules. These restricted schedules are suited to constraint-system-based program analysis. The formalization is based upon a flowgraph-based program model with an operational semantics as reference point. notify = peter.lammich@uni-muenster.de [Shivers-CFA] title = Shivers' Control Flow Analysis topic = Computer science/Programming languages/Static analysis author = Joachim Breitner date = 2010-11-16 abstract = In his dissertation, Olin Shivers introduces a concept of control flow graphs for functional languages, provides an algorithm to statically derive a safe approximation of the control flow graph and proves this algorithm correct. In this research project, Shivers' algorithms and proofs are formalized in the HOLCF extension of HOL. notify = mail@joachim-breitner.de, nipkow@in.tum.de [Slicing] title = Towards Certified Slicing author = Daniel Wasserrab date = 2008-09-16 topic = Computer science/Programming languages/Static analysis abstract = Slicing is a widely-used technique with applications in e.g. compiler technology and software security. Thus verification of algorithms in these areas is often based on the correctness of slicing, which should ideally be proven independent of concrete programming languages and with the help of well-known verifying techniques such as proof assistants. As a first step in this direction, this contribution presents a framework for dynamic and static intraprocedural slicing based on control flow and program dependence graphs. Abstracting from concrete syntax we base the framework on a graph representation of the program fulfilling certain structural and well-formedness properties.

The formalization consists of the basic framework (in subdirectory Basic/), the correctness proof for dynamic slicing (in subdirectory Dynamic/), the correctness proof for static intraprocedural slicing (in subdirectory StaticIntra/) and instantiations of the framework with a simple While language (in subdirectory While/) and the sophisticated object-oriented bytecode language of Jinja (in subdirectory JinjaVM/). For more information on the framework, see the TPHOLS 2008 paper by Wasserrab and Lochbihler and the PLAS 2009 paper by Wasserrab et al. notify = [HRB-Slicing] title = Backing up Slicing: Verifying the Interprocedural Two-Phase Horwitz-Reps-Binkley Slicer author = Daniel Wasserrab date = 2009-11-13 topic = Computer science/Programming languages/Static analysis abstract = After verifying dynamic and static interprocedural slicing, we present a modular framework for static interprocedural slicing. To this end, we formalized the standard two-phase slicer from Horwitz, Reps and Binkley (see their TOPLAS 12(1) 1990 paper) together with summary edges as presented by Reps et al. (see FSE 1994). The framework is again modular in the programming language by using an abstract CFG, defined via structural and well-formedness properties. Using a weak simulation between the original and sliced graph, we were able to prove the correctness of static interprocedural slicing. We also instantiate our framework with a simple While language with procedures. This shows that the chosen abstractions are indeed valid. notify = nipkow@in.tum.de [WorkerWrapper] title = The Worker/Wrapper Transformation author = Peter Gammie date = 2009-10-30 topic = Computer science/Programming languages/Transformations abstract = Gill and Hutton formalise the worker/wrapper transformation, building on the work of Launchbury and Peyton-Jones who developed it as a way of changing the type at which a recursive function operates. This development establishes the soundness of the technique and several examples of its use. notify = peteg42@gmail.com, nipkow@in.tum.de [JiveDataStoreModel] title = Jive Data and Store Model author = Nicole Rauch , Norbert Schirmer <> date = 2005-06-20 license = LGPL topic = Computer science/Programming languages/Misc abstract = This document presents the formalization of an object-oriented data and store model in Isabelle/HOL. This model is being used in the Java Interactive Verification Environment, Jive. notify = kleing@cse.unsw.edu.au, schirmer@in.tum.de [HotelKeyCards] title = Hotel Key Card System author = Tobias Nipkow date = 2006-09-09 topic = Computer science/Security abstract = Two models of an electronic hotel key card system are contrasted: a state based and a trace based one. Both are defined, verified, and proved equivalent in the theorem prover Isabelle/HOL. It is shown that if a guest follows a certain safety policy regarding her key cards, she can be sure that nobody but her can enter her room. notify = nipkow@in.tum.de [RSAPSS] title = SHA1, RSA, PSS and more author = Christina Lindenberg <>, Kai Wirt <> date = 2005-05-02 topic = Computer science/Security/Cryptography abstract = Formal verification is getting more and more important in computer science. However the state of the art formal verification methods in cryptography are very rudimentary. These theories are one step to provide a tool box allowing the use of formal methods in every aspect of cryptography. Moreover we present a proof of concept for the feasibility of verification techniques to a standard signature algorithm. notify = nipkow@in.tum.de [InformationFlowSlicing] title = Information Flow Noninterference via Slicing author = Daniel Wasserrab date = 2010-03-23 topic = Computer science/Security abstract =

In this contribution, we show how correctness proofs for intra- and interprocedural slicing can be used to prove that slicing is able to guarantee information flow noninterference. Moreover, we also illustrate how to lift the control flow graphs of the respective frameworks such that they fulfil the additional assumptions needed in the noninterference proofs. A detailed description of the intraprocedural proof and its interplay with the slicing framework can be found in the PLAS'09 paper by Wasserrab et al.

This entry contains the part for intra-procedural slicing. See entry InformationFlowSlicing_Inter for the inter-procedural part.

extra-history = Change history: [2016-06-10]: The original entry InformationFlowSlicing contained both the inter- and intra-procedural case was split into two for easier maintenance. notify = [InformationFlowSlicing_Inter] title = Inter-Procedural Information Flow Noninterference via Slicing author = Daniel Wasserrab date = 2010-03-23 topic = Computer science/Security abstract =

In this contribution, we show how correctness proofs for intra- and interprocedural slicing can be used to prove that slicing is able to guarantee information flow noninterference. Moreover, we also illustrate how to lift the control flow graphs of the respective frameworks such that they fulfil the additional assumptions needed in the noninterference proofs. A detailed description of the intraprocedural proof and its interplay with the slicing framework can be found in the PLAS'09 paper by Wasserrab et al.

This entry contains the part for inter-procedural slicing. See entry InformationFlowSlicing for the intra-procedural part.

extra-history = Change history: [2016-06-10]: The original entry InformationFlowSlicing contained both the inter- and intra-procedural case was split into two for easier maintenance. notify = [ComponentDependencies] title = Formalisation and Analysis of Component Dependencies author = Maria Spichkova date = 2014-04-28 topic = Computer science/System description languages abstract = This set of theories presents a formalisation in Isabelle/HOL of data dependencies between components. The approach allows to analyse system structure oriented towards efficient checking of system: it aims at elaborating for a concrete system, which parts of the system are necessary to check a given property. notify = maria.spichkova@rmit.edu.au [Verified-Prover] title = A Mechanically Verified, Efficient, Sound and Complete Theorem Prover For First Order Logic author = Tom Ridge <> date = 2004-09-28 topic = Logic/General logic/Mechanization of proofs abstract = Soundness and completeness for a system of first order logic are formally proved, building on James Margetson's formalization of work by Wainer and Wallen. The completeness proofs naturally suggest an algorithm to derive proofs. This algorithm, which can be implemented tail recursively, is formalized in Isabelle/HOL. The algorithm can be executed via the rewriting tactics of Isabelle. Alternatively, the definitions can be exported to OCaml, yielding a directly executable program. notify = lp15@cam.ac.uk [Completeness] title = Completeness theorem author = James Margetson <>, Tom Ridge <> date = 2004-09-20 topic = Logic/Proof theory abstract = The completeness of first-order logic is proved, following the first five pages of Wainer and Wallen's chapter of the book Proof Theory by Aczel et al., CUP, 1992. Their presentation of formulas allows the proofs to use symmetry arguments. Margetson formalized this theorem by early 2000. The Isar conversion is thanks to Tom Ridge. A paper describing the formalization is available [pdf]. notify = lp15@cam.ac.uk [Ordinal] title = Countable Ordinals author = Brian Huffman date = 2005-11-11 topic = Logic/Set theory abstract = This development defines a well-ordered type of countable ordinals. It includes notions of continuous and normal functions, recursively defined functions over ordinals, least fixed-points, and derivatives. Much of ordinal arithmetic is formalized, including exponentials and logarithms. The development concludes with formalizations of Cantor Normal Form and Veblen hierarchies over normal functions. notify = lcp@cl.cam.ac.uk [Ordinals_and_Cardinals] title = Ordinals and Cardinals author = Andrei Popescu date = 2009-09-01 topic = Logic/Set theory abstract = We develop a basic theory of ordinals and cardinals in Isabelle/HOL, up to the point where some cardinality facts relevant for the ``working mathematician" become available. Unlike in set theory, here we do not have at hand canonical notions of ordinal and cardinal. Therefore, here an ordinal is merely a well-order relation and a cardinal is an ordinal minim w.r.t. order embedding on its field. extra-history = Change history: [2012-09-25]: This entry has been discontinued because it is now part of the Isabelle distribution. notify = uuomul@yahoo.com, nipkow@in.tum.de [FOL-Fitting] title = First-Order Logic According to Fitting author = Stefan Berghofer contributors = Asta Halkjær From date = 2007-08-02 topic = Logic/General logic/Classical first-order logic abstract = We present a formalization of parts of Melvin Fitting's book "First-Order Logic and Automated Theorem Proving". The formalization covers the syntax of first-order logic, its semantics, the model existence theorem, a natural deduction proof calculus together with a proof of correctness and completeness, as well as the Löwenheim-Skolem theorem. extra-history = Change history: [2018-07-21]: Proved completeness theorem for open formulas. Proofs are now written in the declarative style. Enumeration of pairs and datatypes is automated using the Countable theory. notify = berghofe@in.tum.de [Epistemic_Logic] title = Epistemic Logic: Completeness of Modal Logics author = Asta Halkjær From topic = Logic/General logic/Logics of knowledge and belief date = 2018-10-29 notify = ahfrom@dtu.dk abstract = This work is a formalization of epistemic logic with countably many agents. It includes proofs of soundness and completeness for the axiom system K. The completeness proof is based on the textbook "Reasoning About Knowledge" by Fagin, Halpern, Moses and Vardi (MIT Press 1995). The extensions of system K (T, KB, K4, S4, S5) and their completeness proofs are based on the textbook "Modal Logic" by Blackburn, de Rijke and Venema (Cambridge University Press 2001). extra-history = Change history: [2021-04-15]: Added completeness of modal logics T, KB, K4, S4 and S5. [SequentInvertibility] title = Invertibility in Sequent Calculi author = Peter Chapman <> date = 2009-08-28 topic = Logic/Proof theory license = LGPL abstract = The invertibility of the rules of a sequent calculus is important for guiding proof search and can be used in some formalised proofs of Cut admissibility. We present sufficient conditions for when a rule is invertible with respect to a calculus. We illustrate the conditions with examples. It must be noted we give purely syntactic criteria; no guarantees are given as to the suitability of the rules. notify = pc@cs.st-andrews.ac.uk, nipkow@in.tum.de [LinearQuantifierElim] title = Quantifier Elimination for Linear Arithmetic author = Tobias Nipkow date = 2008-01-11 topic = Logic/General logic/Decidability of theories abstract = This article formalizes quantifier elimination procedures for dense linear orders, linear real arithmetic and Presburger arithmetic. In each case both a DNF-based non-elementary algorithm and one or more (doubly) exponential NNF-based algorithms are formalized, including the well-known algorithms by Ferrante and Rackoff and by Cooper. The NNF-based algorithms for dense linear orders are new but based on Ferrante and Rackoff and on an algorithm by Loos and Weisspfenning which simulates infenitesimals. All algorithms are directly executable. In particular, they yield reflective quantifier elimination procedures for HOL itself. The formalization makes heavy use of locales and is therefore highly modular. notify = nipkow@in.tum.de [Nat-Interval-Logic] title = Interval Temporal Logic on Natural Numbers author = David Trachtenherz <> date = 2011-02-23 topic = Logic/General logic/Temporal logic abstract = We introduce a theory of temporal logic operators using sets of natural numbers as time domain, formalized in a shallow embedding manner. The theory comprises special natural intervals (theory IL_Interval: open and closed intervals, continuous and modulo intervals, interval traversing results), operators for shifting intervals to left/right on the number axis as well as expanding/contracting intervals by constant factors (theory IL_IntervalOperators.thy), and ultimately definitions and results for unary and binary temporal operators on arbitrary natural sets (theory IL_TemporalOperators). notify = nipkow@in.tum.de [Recursion-Theory-I] title = Recursion Theory I author = Michael Nedzelsky <> date = 2008-04-05 topic = Logic/Computability abstract = This document presents the formalization of introductory material from recursion theory --- definitions and basic properties of primitive recursive functions, Cantor pairing function and computably enumerable sets (including a proof of existence of a one-complete computably enumerable set and a proof of the Rice's theorem). notify = MichaelNedzelsky@yandex.ru [Free-Boolean-Algebra] topic = Logic/General logic/Classical propositional logic title = Free Boolean Algebra author = Brian Huffman date = 2010-03-29 abstract = This theory defines a type constructor representing the free Boolean algebra over a set of generators. Values of type (α)formula represent propositional formulas with uninterpreted variables from type α, ordered by implication. In addition to all the standard Boolean algebra operations, the library also provides a function for building homomorphisms to any other Boolean algebra type. notify = brianh@cs.pdx.edu [Sort_Encodings] title = Sound and Complete Sort Encodings for First-Order Logic author = Jasmin Christian Blanchette , Andrei Popescu date = 2013-06-27 topic = Logic/General logic/Mechanization of proofs abstract = This is a formalization of the soundness and completeness properties for various efficient encodings of sorts in unsorted first-order logic used by Isabelle's Sledgehammer tool.

Essentially, the encodings proceed as follows: a many-sorted problem is decorated with (as few as possible) tags or guards that make the problem monotonic; then sorts can be soundly erased.

The development employs a formalization of many-sorted first-order logic in clausal form (clauses, structures and the basic properties of the satisfaction relation), which could be of interest as the starting point for other formalizations of first-order logic metatheory. notify = uuomul@yahoo.com [Lambda_Free_RPOs] title = Formalization of Recursive Path Orders for Lambda-Free Higher-Order Terms author = Jasmin Christian Blanchette , Uwe Waldmann , Daniel Wand date = 2016-09-23 topic = Logic/Rewriting abstract = This Isabelle/HOL formalization defines recursive path orders (RPOs) for higher-order terms without lambda-abstraction and proves many useful properties about them. The main order fully coincides with the standard RPO on first-order terms also in the presence of currying, distinguishing it from previous work. An optimized variant is formalized as well. It appears promising as the basis of a higher-order superposition calculus. notify = jasmin.blanchette@gmail.com [Lambda_Free_KBOs] title = Formalization of Knuth–Bendix Orders for Lambda-Free Higher-Order Terms author = Heiko Becker , Jasmin Christian Blanchette , Uwe Waldmann , Daniel Wand date = 2016-11-12 topic = Logic/Rewriting abstract = This Isabelle/HOL formalization defines Knuth–Bendix orders for higher-order terms without lambda-abstraction and proves many useful properties about them. The main order fully coincides with the standard transfinite KBO with subterm coefficients on first-order terms. It appears promising as the basis of a higher-order superposition calculus. notify = jasmin.blanchette@gmail.com [Lambda_Free_EPO] title = Formalization of the Embedding Path Order for Lambda-Free Higher-Order Terms author = Alexander Bentkamp topic = Logic/Rewriting date = 2018-10-19 notify = a.bentkamp@vu.nl abstract = This Isabelle/HOL formalization defines the Embedding Path Order (EPO) for higher-order terms without lambda-abstraction and proves many useful properties about it. In contrast to the lambda-free recursive path orders, it does not fully coincide with RPO on first-order terms, but it is compatible with arbitrary higher-order contexts. [Nested_Multisets_Ordinals] title = Formalization of Nested Multisets, Hereditary Multisets, and Syntactic Ordinals author = Jasmin Christian Blanchette , Mathias Fleury , Dmitriy Traytel date = 2016-11-12 topic = Logic/Rewriting abstract = This Isabelle/HOL formalization introduces a nested multiset datatype and defines Dershowitz and Manna's nested multiset order. The order is proved well founded and linear. By removing one constructor, we transform the nested multisets into hereditary multisets. These are isomorphic to the syntactic ordinals—the ordinals can be recursively expressed in Cantor normal form. Addition, subtraction, multiplication, and linear orders are provided on this type. notify = jasmin.blanchette@gmail.com [Abstract-Rewriting] title = Abstract Rewriting topic = Logic/Rewriting date = 2010-06-14 author = Christian Sternagel , René Thiemann license = LGPL abstract = We present an Isabelle formalization of abstract rewriting (see, e.g., the book by Baader and Nipkow). First, we define standard relations like joinability, meetability, conversion, etc. Then, we formalize important properties of abstract rewrite systems, e.g., confluence and strong normalization. Our main concern is on strong normalization, since this formalization is the basis of CeTA (which is mainly about strong normalization of term rewrite systems). Hence lemmas involving strong normalization constitute by far the biggest part of this theory. One of those is Newman's lemma. extra-history = Change history: [2010-09-17]: Added theories defining several (ordered) semirings related to strong normalization and giving some standard instances.
[2013-10-16]: Generalized delta-orders from rationals to Archimedean fields. notify = christian.sternagel@uibk.ac.at, rene.thiemann@uibk.ac.at [First_Order_Terms] title = First-Order Terms author = Christian Sternagel , René Thiemann topic = Logic/Rewriting, Computer science/Algorithms license = LGPL date = 2018-02-06 notify = c.sternagel@gmail.com, rene.thiemann@uibk.ac.at abstract = We formalize basic results on first-order terms, including matching and a first-order unification algorithm, as well as well-foundedness of the subsumption order. This entry is part of the Isabelle Formalization of Rewriting IsaFoR, where first-order terms are omni-present: the unification algorithm is used to certify several confluence and termination techniques, like critical-pair computation and dependency graph approximations; and the subsumption order is a crucial ingredient for completion. [Free-Groups] title = Free Groups author = Joachim Breitner date = 2010-06-24 topic = Mathematics/Algebra abstract = Free Groups are, in a sense, the most generic kind of group. They are defined over a set of generators with no additional relations in between them. They play an important role in the definition of group presentations and in other fields. This theory provides the definition of Free Group as the set of fully canceled words in the generators. The universal property is proven, as well as some isomorphisms results about Free Groups. extra-history = Change history: [2011-12-11]: Added the Ping Pong Lemma. notify = [CofGroups] title = An Example of a Cofinitary Group in Isabelle/HOL author = Bart Kastermans date = 2009-08-04 topic = Mathematics/Algebra abstract = We formalize the usual proof that the group generated by the function k -> k + 1 on the integers gives rise to a cofinitary group. notify = nipkow@in.tum.de [Finitely_Generated_Abelian_Groups] title = Finitely Generated Abelian Groups author = Joseph Thommes<>, Manuel Eberl topic = Mathematics/Algebra date = 2021-07-07 notify = joseph-thommes@gmx.de, manuel@pruvisto.org abstract = This article deals with the formalisation of some group-theoretic results including the fundamental theorem of finitely generated abelian groups characterising the structure of these groups as a uniquely determined product of cyclic groups. Both the invariant factor decomposition and the primary decomposition are covered. Additional work includes results about the direct product, the internal direct product and more group-theoretic lemmas. [Group-Ring-Module] title = Groups, Rings and Modules author = Hidetsune Kobayashi <>, L. Chen <>, H. Murao <> date = 2004-05-18 topic = Mathematics/Algebra abstract = The theory of groups, rings and modules is developed to a great depth. Group theory results include Zassenhaus's theorem and the Jordan-Hoelder theorem. The ring theory development includes ideals, quotient rings and the Chinese remainder theorem. The module development includes the Nakayama lemma, exact sequences and Tensor products. notify = lp15@cam.ac.uk [Robbins-Conjecture] title = A Complete Proof of the Robbins Conjecture author = Matthew Wampler-Doty <> date = 2010-05-22 topic = Mathematics/Algebra abstract = This document gives a formalization of the proof of the Robbins conjecture, following A. Mann, A Complete Proof of the Robbins Conjecture, 2003. notify = nipkow@in.tum.de [Valuation] title = Fundamental Properties of Valuation Theory and Hensel's Lemma author = Hidetsune Kobayashi <> date = 2007-08-08 topic = Mathematics/Algebra abstract = Convergence with respect to a valuation is discussed as convergence of a Cauchy sequence. Cauchy sequences of polynomials are defined. They are used to formalize Hensel's lemma. notify = lp15@cam.ac.uk [Rank_Nullity_Theorem] title = Rank-Nullity Theorem in Linear Algebra author = Jose Divasón , Jesús Aransay topic = Mathematics/Algebra date = 2013-01-16 abstract = In this contribution, we present some formalizations based on the HOL-Multivariate-Analysis session of Isabelle. Firstly, a generalization of several theorems of such library are presented. Secondly, some definitions and proofs involving Linear Algebra and the four fundamental subspaces of a matrix are shown. Finally, we present a proof of the result known in Linear Algebra as the ``Rank-Nullity Theorem'', which states that, given any linear map f from a finite dimensional vector space V to a vector space W, then the dimension of V is equal to the dimension of the kernel of f (which is a subspace of V) and the dimension of the range of f (which is a subspace of W). The proof presented here is based on the one given by Sheldon Axler in his book Linear Algebra Done Right. As a corollary of the previous theorem, and taking advantage of the relationship between linear maps and matrices, we prove that, for every matrix A (which has associated a linear map between finite dimensional vector spaces), the sum of its null space and its column space (which is equal to the range of the linear map) is equal to the number of columns of A. extra-history = Change history: [2014-07-14]: Added some generalizations that allow us to formalize the Rank-Nullity Theorem over finite dimensional vector spaces, instead of over the more particular euclidean spaces. Updated abstract. notify = jose.divasonm@unirioja.es, jesus-maria.aransay@unirioja.es [Affine_Arithmetic] title = Affine Arithmetic author = Fabian Immler date = 2014-02-07 topic = Mathematics/Analysis abstract = We give a formalization of affine forms as abstract representations of zonotopes. We provide affine operations as well as overapproximations of some non-affine operations like multiplication and division. Expressions involving those operations can automatically be turned into (executable) functions approximating the original expression in affine arithmetic. extra-history = Change history: [2015-01-31]: added algorithm for zonotope/hyperplane intersection
[2017-09-20]: linear approximations for all symbols from the floatarith data type notify = immler@in.tum.de [Laplace_Transform] title = Laplace Transform author = Fabian Immler topic = Mathematics/Analysis date = 2019-08-14 notify = fimmler@cs.cmu.edu abstract = This entry formalizes the Laplace transform and concrete Laplace transforms for arithmetic functions, frequency shift, integration and (higher) differentiation in the time domain. It proves Lerch's lemma and uniqueness of the Laplace transform for continuous functions. In order to formalize the foundational assumptions, this entry contains a formalization of piecewise continuous functions and functions of exponential order. [Cauchy] title = Cauchy's Mean Theorem and the Cauchy-Schwarz Inequality author = Benjamin Porter <> date = 2006-03-14 topic = Mathematics/Analysis abstract = This document presents the mechanised proofs of two popular theorems attributed to Augustin Louis Cauchy - Cauchy's Mean Theorem and the Cauchy-Schwarz Inequality. notify = kleing@cse.unsw.edu.au [Integration] title = Integration theory and random variables author = Stefan Richter date = 2004-11-19 topic = Mathematics/Analysis abstract = Lebesgue-style integration plays a major role in advanced probability. We formalize concepts of elementary measure theory, real-valued random variables as Borel-measurable functions, and a stepwise inductive definition of the integral itself. All proofs are carried out in human readable style using the Isar language. extra-note = Note: This article is of historical interest only. Lebesgue-style integration and probability theory are now available as part of the Isabelle/HOL distribution (directory Probability). notify = richter@informatik.rwth-aachen.de, nipkow@in.tum.de, hoelzl@in.tum.de [Ordinary_Differential_Equations] title = Ordinary Differential Equations author = Fabian Immler , Johannes Hölzl topic = Mathematics/Analysis date = 2012-04-26 abstract =

Session Ordinary-Differential-Equations formalizes ordinary differential equations (ODEs) and initial value problems. This work comprises proofs for local and global existence of unique solutions (Picard-Lindelöf theorem). Moreover, it contains a formalization of the (continuous or even differentiable) dependency of the flow on initial conditions as the flow of ODEs.

Not in the generated document are the following sessions:

  • HOL-ODE-Numerics: Rigorous numerical algorithms for computing enclosures of solutions based on Runge-Kutta methods and affine arithmetic. Reachability analysis with splitting and reduction at hyperplanes.
  • HOL-ODE-Examples: Applications of the numerical algorithms to concrete systems of ODEs.
  • Lorenz_C0, Lorenz_C1: Verified algorithms for checking C1-information according to Tucker's proof, computation of C0-information.

extra-history = Change history: [2014-02-13]: added an implementation of the Euler method based on affine arithmetic
[2016-04-14]: added flow and variational equation
[2016-08-03]: numerical algorithms for reachability analysis (using second-order Runge-Kutta methods, splitting, and reduction) implemented using Lammich's framework for automatic refinement
[2017-09-20]: added Poincare map and propagation of variational equation in reachability analysis, verified algorithms for C1-information and computations for C0-information of the Lorenz attractor. notify = immler@in.tum.de, hoelzl@in.tum.de [Polynomials] title = Executable Multivariate Polynomials author = Christian Sternagel , René Thiemann , Alexander Maletzky , Fabian Immler , Florian Haftmann , Andreas Lochbihler , Alexander Bentkamp date = 2010-08-10 topic = Mathematics/Analysis, Mathematics/Algebra, Computer science/Algorithms/Mathematical license = LGPL abstract = We define multivariate polynomials over arbitrary (ordered) semirings in combination with (executable) operations like addition, multiplication, and substitution. We also define (weak) monotonicity of polynomials and comparison of polynomials where we provide standard estimations like absolute positiveness or the more recent approach of Neurauter, Zankl, and Middeldorp. Moreover, it is proven that strongly normalizing (monotone) orders can be lifted to strongly normalizing (monotone) orders over polynomials. Our formalization was performed as part of the IsaFoR/CeTA-system which contains several termination techniques. The provided theories have been essential to formalize polynomial interpretations.

This formalization also contains an abstract representation as coefficient functions with finite support and a type of power-products. If this type is ordered by a linear (term) ordering, various additional notions, such as leading power-product, leading coefficient etc., are introduced as well. Furthermore, a lot of generic properties of, and functions on, multivariate polynomials are formalized, including the substitution and evaluation homomorphisms, embeddings of polynomial rings into larger rings (i.e. with one additional indeterminate), homogenization and dehomogenization of polynomials, and the canonical isomorphism between R[X,Y] and R[X][Y]. extra-history = Change history: [2010-09-17]: Moved theories on arbitrary (ordered) semirings to Abstract Rewriting.
[2016-10-28]: Added abstract representation of polynomials and authors Maletzky/Immler.
[2018-01-23]: Added authors Haftmann, Lochbihler after incorporating their formalization of multivariate polynomials based on Polynomial mappings. Moved material from Bentkamp's entry "Deep Learning".
[2019-04-18]: Added material about polynomials whose power-products are represented themselves by polynomial mappings. notify = rene.thiemann@uibk.ac.at, christian.sternagel@uibk.ac.at, alexander.maletzky@risc.jku.at, immler@in.tum.de [Sqrt_Babylonian] title = Computing N-th Roots using the Babylonian Method author = René Thiemann date = 2013-01-03 topic = Mathematics/Analysis license = LGPL abstract = We implement the Babylonian method to compute n-th roots of numbers. We provide precise algorithms for naturals, integers and rationals, and offer an approximation algorithm for square roots over linear ordered fields. Moreover, there are precise algorithms to compute the floor and the ceiling of n-th roots. extra-history = Change history: [2013-10-16]: Added algorithms to compute floor and ceiling of sqrt of integers. [2014-07-11]: Moved NthRoot_Impl from Real-Impl to this entry. notify = rene.thiemann@uibk.ac.at [Sturm_Sequences] title = Sturm's Theorem author = Manuel Eberl date = 2014-01-11 topic = Mathematics/Analysis abstract = Sturm's Theorem states that polynomial sequences with certain properties, so-called Sturm sequences, can be used to count the number of real roots of a real polynomial. This work contains a proof of Sturm's Theorem and code for constructing Sturm sequences efficiently. It also provides the “sturm” proof method, which can decide certain statements about the roots of real polynomials, such as “the polynomial P has exactly n roots in the interval I” or “P(x) > Q(x) for all x ∈ ℝ”. notify = manuel@pruvisto.org [Sturm_Tarski] title = The Sturm-Tarski Theorem author = Wenda Li date = 2014-09-19 topic = Mathematics/Analysis abstract = We have formalized the Sturm-Tarski theorem (also referred as the Tarski theorem), which generalizes Sturm's theorem. Sturm's theorem is usually used as a way to count distinct real roots, while the Sturm-Tarksi theorem forms the basis for Tarski's classic quantifier elimination for real closed field. notify = wl302@cam.ac.uk [Markov_Models] title = Markov Models author = Johannes Hölzl , Tobias Nipkow date = 2012-01-03 topic = Mathematics/Probability theory, Computer science/Automata and formal languages abstract = This is a formalization of Markov models in Isabelle/HOL. It builds on Isabelle's probability theory. The available models are currently Discrete-Time Markov Chains and a extensions of them with rewards.

As application of these models we formalize probabilistic model checking of pCTL formulas, analysis of IPv4 address allocation in ZeroConf and an analysis of the anonymity of the Crowds protocol. See here for the corresponding paper. notify = hoelzl@in.tum.de [Probabilistic_System_Zoo] title = A Zoo of Probabilistic Systems author = Johannes Hölzl , Andreas Lochbihler , Dmitriy Traytel date = 2015-05-27 topic = Computer science/Automata and formal languages abstract = Numerous models of probabilistic systems are studied in the literature. Coalgebra has been used to classify them into system types and compare their expressiveness. We formalize the resulting hierarchy of probabilistic system types by modeling the semantics of the different systems as codatatypes. This approach yields simple and concise proofs, as bisimilarity coincides with equality for codatatypes.

This work is described in detail in the ITP 2015 publication by the authors. notify = traytel@in.tum.de [Density_Compiler] title = A Verified Compiler for Probability Density Functions author = Manuel Eberl , Johannes Hölzl , Tobias Nipkow date = 2014-10-09 topic = Mathematics/Probability theory, Computer science/Programming languages/Compiling abstract = Bhat et al. [TACAS 2013] developed an inductive compiler that computes density functions for probability spaces described by programs in a probabilistic functional language. In this work, we implement such a compiler for a modified version of this language within the theorem prover Isabelle and give a formal proof of its soundness w.r.t. the semantics of the source and target language. Together with Isabelle's code generation for inductive predicates, this yields a fully verified, executable density compiler. The proof is done in two steps: First, an abstract compiler working with abstract functions modelled directly in the theorem prover's logic is defined and proved sound. Then, this compiler is refined to a concrete version that returns a target-language expression.

An article with the same title and authors is published in the proceedings of ESOP 2015. A detailed presentation of this work can be found in the first author's master's thesis. notify = hoelzl@in.tum.de [CAVA_Automata] title = The CAVA Automata Library author = Peter Lammich date = 2014-05-28 topic = Computer science/Automata and formal languages abstract = We report on the graph and automata library that is used in the fully verified LTL model checker CAVA. As most components of CAVA use some type of graphs or automata, a common automata library simplifies assembly of the components and reduces redundancy.

The CAVA Automata Library provides a hierarchy of graph and automata classes, together with some standard algorithms. Its object oriented design allows for sharing of algorithms, theorems, and implementations between its classes, and also simplifies extensions of the library. Moreover, it is integrated into the Automatic Refinement Framework, supporting automatic refinement of the abstract automata types to efficient data structures.

Note that the CAVA Automata Library is work in progress. Currently, it is very specifically tailored towards the requirements of the CAVA model checker. Nevertheless, the formalization techniques presented here allow an extension of the library to a wider scope. Moreover, they are not limited to graph libraries, but apply to class hierarchies in general.

The CAVA Automata Library is described in the paper: Peter Lammich, The CAVA Automata Library, Isabelle Workshop 2014. notify = lammich@in.tum.de [LTL] title = Linear Temporal Logic author = Salomon Sickert contributors = Benedikt Seidl date = 2016-03-01 topic = Logic/General logic/Temporal logic, Computer science/Automata and formal languages abstract = This theory provides a formalisation of linear temporal logic (LTL) and unifies previous formalisations within the AFP. This entry establishes syntax and semantics for this logic and decouples it from existing entries, yielding a common environment for theories reasoning about LTL. Furthermore a parser written in SML and an executable simplifier are provided. extra-history = Change history: [2019-03-12]: Support for additional operators, implementation of common equivalence relations, definition of syntactic fragments of LTL and the minimal disjunctive normal form.
notify = sickert@in.tum.de [LTL_to_GBA] title = Converting Linear-Time Temporal Logic to Generalized Büchi Automata author = Alexander Schimpf , Peter Lammich date = 2014-05-28 topic = Computer science/Automata and formal languages abstract = We formalize linear-time temporal logic (LTL) and the algorithm by Gerth et al. to convert LTL formulas to generalized Büchi automata. We also formalize some syntactic rewrite rules that can be applied to optimize the LTL formula before conversion. Moreover, we integrate the Stuttering Equivalence AFP-Entry by Stefan Merz, adapting the lemma that next-free LTL formula cannot distinguish between stuttering equivalent runs to our setting.

We use the Isabelle Refinement and Collection framework, as well as the Autoref tool, to obtain a refined version of our algorithm, from which efficiently executable code can be extracted. notify = lammich@in.tum.de [Gabow_SCC] title = Verified Efficient Implementation of Gabow's Strongly Connected Components Algorithm author = Peter Lammich date = 2014-05-28 topic = Computer science/Algorithms/Graph, Mathematics/Graph theory abstract = We present an Isabelle/HOL formalization of Gabow's algorithm for finding the strongly connected components of a directed graph. Using data refinement techniques, we extract efficient code that performs comparable to a reference implementation in Java. Our style of formalization allows for re-using large parts of the proofs when defining variants of the algorithm. We demonstrate this by verifying an algorithm for the emptiness check of generalized Büchi automata, re-using most of the existing proofs. notify = lammich@in.tum.de [Promela] title = Promela Formalization author = René Neumann date = 2014-05-28 topic = Computer science/System description languages abstract = We present an executable formalization of the language Promela, the description language for models of the model checker SPIN. This formalization is part of the work for a completely verified model checker (CAVA), but also serves as a useful (and executable!) description of the semantics of the language itself, something that is currently missing. The formalization uses three steps: It takes an abstract syntax tree generated from an SML parser, removes syntactic sugar and enriches it with type information. This further gets translated into a transition system, on which the semantic engine (read: successor function) operates. notify = [CAVA_LTL_Modelchecker] title = A Fully Verified Executable LTL Model Checker author = Javier Esparza , Peter Lammich , René Neumann , Tobias Nipkow , Alexander Schimpf , Jan-Georg Smaus date = 2014-05-28 topic = Computer science/Automata and formal languages abstract = We present an LTL model checker whose code has been completely verified using the Isabelle theorem prover. The checker consists of over 4000 lines of ML code. The code is produced using the Isabelle Refinement Framework, which allows us to split its correctness proof into (1) the proof of an abstract version of the checker, consisting of a few hundred lines of ``formalized pseudocode'', and (2) a verified refinement step in which mathematical sets and other abstract structures are replaced by implementations of efficient structures like red-black trees and functional arrays. This leads to a checker that, while still slower than unverified checkers, can already be used as a trusted reference implementation against which advanced implementations can be tested.

An early version of this model checker is described in the CAV 2013 paper with the same title. notify = lammich@in.tum.de [Fermat3_4] title = Fermat's Last Theorem for Exponents 3 and 4 and the Parametrisation of Pythagorean Triples author = Roelof Oosterhuis <> date = 2007-08-12 topic = Mathematics/Number theory abstract = This document presents the mechanised proofs of

  • Fermat's Last Theorem for exponents 3 and 4 and
  • the parametrisation of Pythagorean Triples.
notify = nipkow@in.tum.de, roelofoosterhuis@gmail.com [Perfect-Number-Thm] title = Perfect Number Theorem author = Mark Ijbema date = 2009-11-22 topic = Mathematics/Number theory abstract = These theories present the mechanised proof of the Perfect Number Theorem. notify = nipkow@in.tum.de [SumSquares] title = Sums of Two and Four Squares author = Roelof Oosterhuis <> date = 2007-08-12 topic = Mathematics/Number theory abstract = This document presents the mechanised proofs of the following results:
  • any prime number of the form 4m+1 can be written as the sum of two squares;
  • any natural number can be written as the sum of four squares
notify = nipkow@in.tum.de, roelofoosterhuis@gmail.com [Lehmer] title = Lehmer's Theorem author = Simon Wimmer , Lars Noschinski date = 2013-07-22 topic = Mathematics/Number theory abstract = In 1927, Lehmer presented criterions for primality, based on the converse of Fermat's litte theorem. This work formalizes the second criterion from Lehmer's paper, a necessary and sufficient condition for primality.

As a side product we formalize some properties of Euler's phi-function, the notion of the order of an element of a group, and the cyclicity of the multiplicative group of a finite field. notify = noschinl@gmail.com, simon.wimmer@tum.de [Pratt_Certificate] title = Pratt's Primality Certificates author = Simon Wimmer , Lars Noschinski date = 2013-07-22 topic = Mathematics/Number theory abstract = In 1975, Pratt introduced a proof system for certifying primes. He showed that a number p is prime iff a primality certificate for p exists. By showing a logarithmic upper bound on the length of the certificates in size of the prime number, he concluded that the decision problem for prime numbers is in NP. This work formalizes soundness and completeness of Pratt's proof system as well as an upper bound for the size of the certificate. notify = noschinl@gmail.com, simon.wimmer@tum.de [Monad_Memo_DP] title = Monadification, Memoization and Dynamic Programming author = Simon Wimmer , Shuwei Hu , Tobias Nipkow topic = Computer science/Programming languages/Transformations, Computer science/Algorithms, Computer science/Functional programming date = 2018-05-22 notify = wimmers@in.tum.de abstract = We present a lightweight framework for the automatic verified (functional or imperative) memoization of recursive functions. Our tool can turn a pure Isabelle/HOL function definition into a monadified version in a state monad or the Imperative HOL heap monad, and prove a correspondence theorem. We provide a variety of memory implementations for the two types of monads. A number of simple techniques allow us to achieve bottom-up computation and space-efficient memoization. The framework’s utility is demonstrated on a number of representative dynamic programming problems. A detailed description of our work can be found in the accompanying paper [2]. [Probabilistic_Timed_Automata] title = Probabilistic Timed Automata author = Simon Wimmer , Johannes Hölzl topic = Mathematics/Probability theory, Computer science/Automata and formal languages date = 2018-05-24 notify = wimmers@in.tum.de, hoelzl@in.tum.de abstract = We present a formalization of probabilistic timed automata (PTA) for which we try to follow the formula MDP + TA = PTA as far as possible: our work starts from our existing formalizations of Markov decision processes (MDP) and timed automata (TA) and combines them modularly. We prove the fundamental result for probabilistic timed automata: the region construction that is known from timed automata carries over to the probabilistic setting. In particular, this allows us to prove that minimum and maximum reachability probabilities can be computed via a reduction to MDP model checking, including the case where one wants to disregard unrealizable behavior. Further information can be found in our ITP paper [2]. [Hidden_Markov_Models] title = Hidden Markov Models author = Simon Wimmer topic = Mathematics/Probability theory, Computer science/Algorithms date = 2018-05-25 notify = wimmers@in.tum.de abstract = This entry contains a formalization of hidden Markov models [3] based on Johannes Hölzl's formalization of discrete time Markov chains [1]. The basic definitions are provided and the correctness of two main (dynamic programming) algorithms for hidden Markov models is proved: the forward algorithm for computing the likelihood of an observed sequence, and the Viterbi algorithm for decoding the most probable hidden state sequence. The Viterbi algorithm is made executable including memoization. Hidden markov models have various applications in natural language processing. For an introduction see Jurafsky and Martin [2]. [ArrowImpossibilityGS] title = Arrow and Gibbard-Satterthwaite author = Tobias Nipkow date = 2008-09-01 topic = Mathematics/Games and economics abstract = This article formalizes two proofs of Arrow's impossibility theorem due to Geanakoplos and derives the Gibbard-Satterthwaite theorem as a corollary. One formalization is based on utility functions, the other one on strict partial orders.

An article about these proofs is found here. notify = nipkow@in.tum.de [SenSocialChoice] title = Some classical results in Social Choice Theory author = Peter Gammie date = 2008-11-09 topic = Mathematics/Games and economics abstract = Drawing on Sen's landmark work "Collective Choice and Social Welfare" (1970), this development proves Arrow's General Possibility Theorem, Sen's Liberal Paradox and May's Theorem in a general setting. The goal was to make precise the classical statements and proofs of these results, and to provide a foundation for more recent results such as the Gibbard-Satterthwaite and Duggan-Schwartz theorems. notify = nipkow@in.tum.de [Vickrey_Clarke_Groves] title = VCG - Combinatorial Vickrey-Clarke-Groves Auctions author = Marco B. Caminati <>, Manfred Kerber , Christoph Lange, Colin Rowat date = 2015-04-30 topic = Mathematics/Games and economics abstract = A VCG auction (named after their inventors Vickrey, Clarke, and Groves) is a generalization of the single-good, second price Vickrey auction to the case of a combinatorial auction (multiple goods, from which any participant can bid on each possible combination). We formalize in this entry VCG auctions, including tie-breaking and prove that the functions for the allocation and the price determination are well-defined. Furthermore we show that the allocation function allocates goods only to participants, only goods in the auction are allocated, and no good is allocated twice. We also show that the price function is non-negative. These properties also hold for the automatically extracted Scala code. notify = mnfrd.krbr@gmail.com [Topology] title = Topology author = Stefan Friedrich <> date = 2004-04-26 topic = Mathematics/Topology abstract = This entry contains two theories. The first, Topology, develops the basic notions of general topology. The second, which can be viewed as a demonstration of the first, is called LList_Topology. It develops the topology of lazy lists. notify = lcp@cl.cam.ac.uk [Knot_Theory] title = Knot Theory author = T.V.H. Prathamesh date = 2016-01-20 topic = Mathematics/Topology abstract = This work contains a formalization of some topics in knot theory. The concepts that were formalized include definitions of tangles, links, framed links and link/tangle equivalence. The formalization is based on a formulation of links in terms of tangles. We further construct and prove the invariance of the Bracket polynomial. Bracket polynomial is an invariant of framed links closely linked to the Jones polynomial. This is perhaps the first attempt to formalize any aspect of knot theory in an interactive proof assistant. notify = prathamesh@imsc.res.in [Graph_Theory] title = Graph Theory author = Lars Noschinski date = 2013-04-28 topic = Mathematics/Graph theory abstract = This development provides a formalization of directed graphs, supporting (labelled) multi-edges and infinite graphs. A polymorphic edge type allows edges to be treated as pairs of vertices, if multi-edges are not required. Formalized properties are i.a. walks (and related concepts), connectedness and subgraphs and basic properties of isomorphisms.

This formalization is used to prove characterizations of Euler Trails, Shortest Paths and Kuratowski subgraphs. notify = noschinl@gmail.com [Planarity_Certificates] title = Planarity Certificates author = Lars Noschinski date = 2015-11-11 topic = Mathematics/Graph theory abstract = This development provides a formalization of planarity based on combinatorial maps and proves that Kuratowski's theorem implies combinatorial planarity. Moreover, it contains verified implementations of programs checking certificates for planarity (i.e., a combinatorial map) or non-planarity (i.e., a Kuratowski subgraph). notify = noschinl@gmail.com [Max-Card-Matching] title = Maximum Cardinality Matching author = Christine Rizkallah date = 2011-07-21 topic = Mathematics/Graph theory abstract =

A matching in a graph G is a subset M of the edges of G such that no two share an endpoint. A matching has maximum cardinality if its cardinality is at least as large as that of any other matching. An odd-set cover OSC of a graph G is a labeling of the nodes of G with integers such that every edge of G is either incident to a node labeled 1 or connects two nodes labeled with the same number i ≥ 2.

This article proves Edmonds theorem:
Let M be a matching in a graph G and let OSC be an odd-set cover of G. For any i ≥ 0, let n(i) be the number of nodes labeled i. If |M| = n(1) + ∑i ≥ 2(n(i) div 2), then M is a maximum cardinality matching.

notify = nipkow@in.tum.de [Girth_Chromatic] title = A Probabilistic Proof of the Girth-Chromatic Number Theorem author = Lars Noschinski date = 2012-02-06 topic = Mathematics/Graph theory abstract = This works presents a formalization of the Girth-Chromatic number theorem in graph theory, stating that graphs with arbitrarily large girth and chromatic number exist. The proof uses the theory of Random Graphs to prove the existence with probabilistic arguments. notify = noschinl@gmail.com [Random_Graph_Subgraph_Threshold] title = Properties of Random Graphs -- Subgraph Containment author = Lars Hupel date = 2014-02-13 topic = Mathematics/Graph theory, Mathematics/Probability theory abstract = Random graphs are graphs with a fixed number of vertices, where each edge is present with a fixed probability. We are interested in the probability that a random graph contains a certain pattern, for example a cycle or a clique. A very high edge probability gives rise to perhaps too many edges (which degrades performance for many algorithms), whereas a low edge probability might result in a disconnected graph. We prove a theorem about a threshold probability such that a higher edge probability will asymptotically almost surely produce a random graph with the desired subgraph. notify = hupel@in.tum.de [Flyspeck-Tame] title = Flyspeck I: Tame Graphs author = Gertrud Bauer <>, Tobias Nipkow date = 2006-05-22 topic = Mathematics/Graph theory abstract = These theories present the verified enumeration of tame plane graphs as defined by Thomas C. Hales in his proof of the Kepler Conjecture in his book Dense Sphere Packings. A Blueprint for Formal Proofs. [CUP 2012]. The values of the constants in the definition of tameness are identical to those in the Flyspeck project. The IJCAR 2006 paper by Nipkow, Bauer and Schultz refers to the original version of Hales' proof, the ITP 2011 paper by Nipkow refers to the Blueprint version of the proof. extra-history = Change history: [2010-11-02]: modified theories to reflect the modified definition of tameness in Hales' revised proof.
[2014-07-03]: modified constants in def of tameness and Archive according to the final state of the Flyspeck proof. notify = nipkow@in.tum.de [Well_Quasi_Orders] title = Well-Quasi-Orders author = Christian Sternagel date = 2012-04-13 topic = Mathematics/Combinatorics abstract = Based on Isabelle/HOL's type class for preorders, we introduce a type class for well-quasi-orders (wqo) which is characterized by the absence of "bad" sequences (our proofs are along the lines of the proof of Nash-Williams, from which we also borrow terminology). Our main results are instantiations for the product type, the list type, and a type of finite trees, which (almost) directly follow from our proofs of (1) Dickson's Lemma, (2) Higman's Lemma, and (3) Kruskal's Tree Theorem. More concretely:
  • If the sets A and B are wqo then their Cartesian product is wqo.
  • If the set A is wqo then the set of finite lists over A is wqo.
  • If the set A is wqo then the set of finite trees over A is wqo.
The research was funded by the Austrian Science Fund (FWF): J3202. extra-history = Change history: [2012-06-11]: Added Kruskal's Tree Theorem.
[2012-12-19]: New variant of Kruskal's tree theorem for terms (as opposed to variadic terms, i.e., trees), plus finite version of the tree theorem as corollary.
[2013-05-16]: Simplified construction of minimal bad sequences.
[2014-07-09]: Simplified proofs of Higman's lemma and Kruskal's tree theorem, based on homogeneous sequences.
[2016-01-03]: An alternative proof of Higman's lemma by open induction.
[2017-06-08]: Proved (classical) equivalence to inductive definition of almost-full relations according to the ITP 2012 paper "Stop When You Are Almost-Full" by Vytiniotis, Coquand, and Wahlstedt. notify = c.sternagel@gmail.com [Marriage] title = Hall's Marriage Theorem author = Dongchen Jiang , Tobias Nipkow date = 2010-12-17 topic = Mathematics/Combinatorics abstract = Two proofs of Hall's Marriage Theorem: one due to Halmos and Vaughan, one due to Rado. extra-history = Change history: [2011-09-09]: Added Rado's proof notify = nipkow@in.tum.de [Bondy] title = Bondy's Theorem author = Jeremy Avigad , Stefan Hetzl date = 2012-10-27 topic = Mathematics/Combinatorics abstract = A proof of Bondy's theorem following B. Bollabas, Combinatorics, 1986, Cambridge University Press. notify = avigad@cmu.edu, hetzl@logic.at [Ramsey-Infinite] title = Ramsey's theorem, infinitary version author = Tom Ridge <> date = 2004-09-20 topic = Mathematics/Combinatorics abstract = This formalization of Ramsey's theorem (infinitary version) is taken from Boolos and Jeffrey, Computability and Logic, 3rd edition, Chapter 26. It differs slightly from the text by assuming a slightly stronger hypothesis. In particular, the induction hypothesis is stronger, holding for any infinite subset of the naturals. This avoids the rather peculiar mapping argument between kj and aikj on p.263, which is unnecessary and slightly mars this really beautiful result. notify = lp15@cam.ac.uk [Derangements] title = Derangements Formula author = Lukas Bulwahn date = 2015-06-27 topic = Mathematics/Combinatorics abstract = The Derangements Formula describes the number of fixpoint-free permutations as a closed formula. This theorem is the 88th theorem in a list of the ``Top 100 Mathematical Theorems''. notify = lukas.bulwahn@gmail.com [Euler_Partition] title = Euler's Partition Theorem author = Lukas Bulwahn date = 2015-11-19 topic = Mathematics/Combinatorics abstract = Euler's Partition Theorem states that the number of partitions with only distinct parts is equal to the number of partitions with only odd parts. The combinatorial proof follows John Harrison's HOL Light formalization. This theorem is the 45th theorem of the Top 100 Theorems list. notify = lukas.bulwahn@gmail.com [Discrete_Summation] title = Discrete Summation author = Florian Haftmann contributors = Amine Chaieb <> date = 2014-04-13 topic = Mathematics/Combinatorics abstract = These theories introduce basic concepts and proofs about discrete summation: shifts, formal summation, falling factorials and stirling numbers. As proof of concept, a simple summation conversion is provided. notify = florian.haftmann@informatik.tu-muenchen.de [Open_Induction] title = Open Induction author = Mizuhito Ogawa <>, Christian Sternagel date = 2012-11-02 topic = Mathematics/Combinatorics abstract = A proof of the open induction schema based on J.-C. Raoult, Proving open properties by induction, Information Processing Letters 29, 1988, pp.19-23.

This research was supported by the Austrian Science Fund (FWF): J3202.

notify = c.sternagel@gmail.com [Category] title = Category Theory to Yoneda's Lemma author = Greg O'Keefe date = 2005-04-21 topic = Mathematics/Category theory license = LGPL abstract = This development proves Yoneda's lemma and aims to be readable by humans. It only defines what is needed for the lemma: categories, functors and natural transformations. Limits, adjunctions and other important concepts are not included. extra-history = Change history: [2010-04-23]: The definition of the constant equinumerous was slightly too weak in the original submission and has been fixed in revision 8c2b5b3c995f. notify = lcp@cl.cam.ac.uk [Category2] title = Category Theory author = Alexander Katovsky date = 2010-06-20 topic = Mathematics/Category theory abstract = This article presents a development of Category Theory in Isabelle/HOL. A Category is defined using records and locales. Functors and Natural Transformations are also defined. The main result that has been formalized is that the Yoneda functor is a full and faithful embedding. We also formalize the completeness of many sorted monadic equational logic. Extensive use is made of the HOLZF theory in both cases. For an informal description see here [pdf]. notify = alexander.katovsky@cantab.net [FunWithFunctions] title = Fun With Functions author = Tobias Nipkow date = 2008-08-26 topic = Mathematics/Misc abstract = This is a collection of cute puzzles of the form ``Show that if a function satisfies the following constraints, it must be ...'' Please add further examples to this collection! notify = nipkow@in.tum.de [FunWithTilings] title = Fun With Tilings author = Tobias Nipkow , Lawrence C. Paulson date = 2008-11-07 topic = Mathematics/Misc abstract = Tilings are defined inductively. It is shown that one form of mutilated chess board cannot be tiled with dominoes, while another one can be tiled with L-shaped tiles. Please add further fun examples of this kind! notify = nipkow@in.tum.de [Lazy-Lists-II] title = Lazy Lists II author = Stefan Friedrich <> date = 2004-04-26 topic = Computer science/Data structures abstract = This theory contains some useful extensions to the LList (lazy list) theory by Larry Paulson, including finite, infinite, and positive llists over an alphabet, as well as the new constants take and drop and the prefix order of llists. Finally, the notions of safety and liveness in the sense of Alpern and Schneider (1985) are defined. notify = lcp@cl.cam.ac.uk [Ribbon_Proofs] title = Ribbon Proofs author = John Wickerson <> date = 2013-01-19 topic = Computer science/Programming languages/Logics abstract = This document concerns the theory of ribbon proofs: a diagrammatic proof system, based on separation logic, for verifying program correctness. We include the syntax, proof rules, and soundness results for two alternative formalisations of ribbon proofs.

Compared to traditional proof outlines, ribbon proofs emphasise the structure of a proof, so are intelligible and pedagogical. Because they contain less redundancy than proof outlines, and allow each proof step to be checked locally, they may be more scalable. Where proof outlines are cumbersome to modify, ribbon proofs can be visually manoeuvred to yield proofs of variant programs. notify = [Koenigsberg_Friendship] title = The Königsberg Bridge Problem and the Friendship Theorem author = Wenda Li date = 2013-07-19 topic = Mathematics/Graph theory abstract = This development provides a formalization of undirected graphs and simple graphs, which are based on Benedikt Nordhoff and Peter Lammich's simple formalization of labelled directed graphs in the archive. Then, with our formalization of graphs, we show both necessary and sufficient conditions for Eulerian trails and circuits as well as the fact that the Königsberg Bridge Problem does not have a solution. In addition, we show the Friendship Theorem in simple graphs. notify = [Tree_Decomposition] title = Tree Decomposition author = Christoph Dittmann notify = date = 2016-05-31 topic = Mathematics/Graph theory abstract = We formalize tree decompositions and tree width in Isabelle/HOL, proving that trees have treewidth 1. We also show that every edge of a tree decomposition is a separation of the underlying graph. As an application of this theorem we prove that complete graphs of size n have treewidth n-1. [Menger] title = Menger's Theorem author = Christoph Dittmann topic = Mathematics/Graph theory date = 2017-02-26 notify = isabelle@christoph-d.de abstract = We present a formalization of Menger's Theorem for directed and undirected graphs in Isabelle/HOL. This well-known result shows that if two non-adjacent distinct vertices u, v in a directed graph have no separator smaller than n, then there exist n internally vertex-disjoint paths from u to v. The version for undirected graphs follows immediately because undirected graphs are a special case of directed graphs. [IEEE_Floating_Point] title = A Formal Model of IEEE Floating Point Arithmetic author = Lei Yu contributors = Fabian Hellauer , Fabian Immler date = 2013-07-27 topic = Computer science/Data structures abstract = This development provides a formal model of IEEE-754 floating-point arithmetic. This formalization, including formal specification of the standard and proofs of important properties of floating-point arithmetic, forms the foundation for verifying programs with floating-point computation. There is also a code generation setup for floats so that we can execute programs using this formalization in functional programming languages. notify = lp15@cam.ac.uk, immler@in.tum.de extra-history = Change history: [2017-09-25]: Added conversions from and to software floating point numbers (by Fabian Hellauer and Fabian Immler).
[2018-02-05]: 'Modernized' representation following the formalization in HOL4: former "float_format" and predicate "is_valid" is now encoded in a type "('e, 'f) float" where 'e and 'f encode the size of exponent and fraction. [Native_Word] title = Native Word author = Andreas Lochbihler contributors = Peter Lammich date = 2013-09-17 topic = Computer science/Data structures abstract = This entry makes machine words and machine arithmetic available for code generation from Isabelle/HOL. It provides a common abstraction that hides the differences between the different target languages. The code generator maps these operations to the APIs of the target languages. Apart from that, we extend the available bit operations on types int and integer, and map them to the operations in the target languages. extra-history = Change history: [2013-11-06]: added conversion function between native words and characters (revision fd23d9a7fe3a)
[2014-03-31]: added words of default size in the target language (by Peter Lammich) (revision 25caf5065833)
[2014-10-06]: proper test setup with compilation and execution of tests in all target languages (revision 5d7a1c9ae047)
[2017-09-02]: added 64-bit words (revision c89f86244e3c)
[2018-07-15]: added cast operators for default-size words (revision fc1f1fb8dd30)
notify = mail@andreas-lochbihler.de [XML] title = XML author = Christian Sternagel , René Thiemann date = 2014-10-03 topic = Computer science/Functional programming, Computer science/Data structures abstract = This entry provides an XML library for Isabelle/HOL. This includes parsing and pretty printing of XML trees as well as combinators for transforming XML trees into arbitrary user-defined data. The main contribution of this entry is an interface (fit for code generation) that allows for communication between verified programs formalized in Isabelle/HOL and the outside world via XML. This library was developed as part of the IsaFoR/CeTA project to which we refer for examples of its usage. notify = c.sternagel@gmail.com, rene.thiemann@uibk.ac.at [HereditarilyFinite] title = The Hereditarily Finite Sets author = Lawrence C. Paulson date = 2013-11-17 topic = Logic/Set theory abstract = The theory of hereditarily finite sets is formalised, following the development of Swierczkowski. An HF set is a finite collection of other HF sets; they enjoy an induction principle and satisfy all the axioms of ZF set theory apart from the axiom of infinity, which is negated. All constructions that are possible in ZF set theory (Cartesian products, disjoint sums, natural numbers, functions) without using infinite sets are possible here. The definition of addition for the HF sets follows Kirby. This development forms the foundation for the Isabelle proof of Gödel's incompleteness theorems, which has been formalised separately. extra-history = Change history: [2015-02-23]: Added the theory "Finitary" defining the class of types that can be embedded in hf, including int, char, option, list, etc. notify = lp15@cam.ac.uk [Incompleteness] title = Gödel's Incompleteness Theorems author = Lawrence C. Paulson date = 2013-11-17 topic = Logic/Proof theory abstract = Gödel's two incompleteness theorems are formalised, following a careful presentation by Swierczkowski, in the theory of hereditarily finite sets. This represents the first ever machine-assisted proof of the second incompleteness theorem. Compared with traditional formalisations using Peano arithmetic (see e.g. Boolos), coding is simpler, with no need to formalise the notion of multiplication (let alone that of a prime number) in the formalised calculus upon which the theorem is based. However, other technical problems had to be solved in order to complete the argument. notify = lp15@cam.ac.uk [Finite_Automata_HF] title = Finite Automata in Hereditarily Finite Set Theory author = Lawrence C. Paulson date = 2015-02-05 topic = Computer science/Automata and formal languages abstract = Finite Automata, both deterministic and non-deterministic, for regular languages. The Myhill-Nerode Theorem. Closure under intersection, concatenation, etc. Regular expressions define regular languages. Closure under reversal; the powerset construction mapping NFAs to DFAs. Left and right languages; minimal DFAs. Brzozowski's minimization algorithm. Uniqueness up to isomorphism of minimal DFAs. notify = lp15@cam.ac.uk [Decreasing-Diagrams] title = Decreasing Diagrams author = Harald Zankl license = LGPL date = 2013-11-01 topic = Logic/Rewriting abstract = This theory contains a formalization of decreasing diagrams showing that any locally decreasing abstract rewrite system is confluent. We consider the valley (van Oostrom, TCS 1994) and the conversion version (van Oostrom, RTA 2008) and closely follow the original proofs. As an application we prove Newman's lemma. notify = Harald.Zankl@uibk.ac.at [Decreasing-Diagrams-II] title = Decreasing Diagrams II author = Bertram Felgenhauer license = LGPL date = 2015-08-20 topic = Logic/Rewriting abstract = This theory formalizes the commutation version of decreasing diagrams for Church-Rosser modulo. The proof follows Felgenhauer and van Oostrom (RTA 2013). The theory also provides important specializations, in particular van Oostrom’s conversion version (TCS 2008) of decreasing diagrams. notify = bertram.felgenhauer@uibk.ac.at [GoedelGod] title = Gödel's God in Isabelle/HOL author = Christoph Benzmüller , Bruno Woltzenlogel Paleo date = 2013-11-12 topic = Logic/Philosophical aspects abstract = Dana Scott's version of Gödel's proof of God's existence is formalized in quantified modal logic KB (QML KB). QML KB is modeled as a fragment of classical higher-order logic (HOL); thus, the formalization is essentially a formalization in HOL. notify = lp15@cam.ac.uk, c.benzmueller@fu-berlin.de [Types_Tableaus_and_Goedels_God] title = Types, Tableaus and Gödel’s God in Isabelle/HOL author = David Fuenmayor , Christoph Benzmüller topic = Logic/Philosophical aspects date = 2017-05-01 notify = davfuenmayor@gmail.com, c.benzmueller@gmail.com abstract = A computer-formalisation of the essential parts of Fitting's textbook "Types, Tableaus and Gödel's God" in Isabelle/HOL is presented. In particular, Fitting's (and Anderson's) variant of the ontological argument is verified and confirmed. This variant avoids the modal collapse, which has been criticised as an undesirable side-effect of Kurt Gödel's (and Dana Scott's) versions of the ontological argument. Fitting's work is employing an intensional higher-order modal logic, which we shallowly embed here in classical higher-order logic. We then utilize the embedded logic for the formalisation of Fitting's argument. (See also the earlier AFP entry ``Gödel's God in Isabelle/HOL''.) [GewirthPGCProof] title = Formalisation and Evaluation of Alan Gewirth's Proof for the Principle of Generic Consistency in Isabelle/HOL author = David Fuenmayor , Christoph Benzmüller topic = Logic/Philosophical aspects date = 2018-10-30 notify = davfuenmayor@gmail.com, c.benzmueller@gmail.com abstract = An ambitious ethical theory ---Alan Gewirth's "Principle of Generic Consistency"--- is encoded and analysed in Isabelle/HOL. Gewirth's theory has stirred much attention in philosophy and ethics and has been proposed as a potential means to bound the impact of artificial general intelligence. extra-history = Change history: [2019-04-09]: added proof for a stronger variant of the PGC and examplary inferences (revision 88182cb0a2f6)
[Lowe_Ontological_Argument] title = Computer-assisted Reconstruction and Assessment of E. J. Lowe's Modal Ontological Argument author = David Fuenmayor , Christoph Benzmüller topic = Logic/Philosophical aspects date = 2017-09-21 notify = davfuenmayor@gmail.com, c.benzmueller@gmail.com abstract = Computers may help us to understand --not just verify-- philosophical arguments. By utilizing modern proof assistants in an iterative interpretive process, we can reconstruct and assess an argument by fully formal means. Through the mechanization of a variant of St. Anselm's ontological argument by E. J. Lowe, which is a paradigmatic example of a natural-language argument with strong ties to metaphysics and religion, we offer an ideal showcase for our computer-assisted interpretive method. [AnselmGod] title = Anselm's God in Isabelle/HOL author = Ben Blumson topic = Logic/Philosophical aspects date = 2017-09-06 notify = benblumson@gmail.com abstract = Paul Oppenheimer and Edward Zalta's formalisation of Anselm's ontological argument for the existence of God is automated by embedding a free logic for definite descriptions within Isabelle/HOL. [Tail_Recursive_Functions] title = A General Method for the Proof of Theorems on Tail-recursive Functions author = Pasquale Noce date = 2013-12-01 topic = Computer science/Functional programming abstract =

Tail-recursive function definitions are sometimes more straightforward than alternatives, but proving theorems on them may be roundabout because of the peculiar form of the resulting recursion induction rules.

This paper describes a proof method that provides a general solution to this problem by means of suitable invariants over inductive sets, and illustrates the application of such method by examining two case studies.

notify = pasquale.noce.lavoro@gmail.com [CryptoBasedCompositionalProperties] title = Compositional Properties of Crypto-Based Components author = Maria Spichkova date = 2014-01-11 topic = Computer science/Security abstract = This paper presents an Isabelle/HOL set of theories which allows the specification of crypto-based components and the verification of their composition properties wrt. cryptographic aspects. We introduce a formalisation of the security property of data secrecy, the corresponding definitions and proofs. Please note that here we import the Isabelle/HOL theory ListExtras.thy, presented in the AFP entry FocusStreamsCaseStudies-AFP. notify = maria.spichkova@rmit.edu.au [Featherweight_OCL] title = Featherweight OCL: A Proposal for a Machine-Checked Formal Semantics for OCL 2.5 author = Achim D. Brucker , Frédéric Tuong , Burkhart Wolff date = 2014-01-16 topic = Computer science/System description languages abstract = The Unified Modeling Language (UML) is one of the few modeling languages that is widely used in industry. While UML is mostly known as diagrammatic modeling language (e.g., visualizing class models), it is complemented by a textual language, called Object Constraint Language (OCL). The current version of OCL is based on a four-valued logic that turns UML into a formal language. Any type comprises the elements "invalid" and "null" which are propagated as strict and non-strict, respectively. Unfortunately, the former semi-formal semantics of this specification language, captured in the "Annex A" of the OCL standard, leads to different interpretations of corner cases. We formalize the core of OCL: denotational definitions, a logical calculus and operational rules that allow for the execution of OCL expressions by a mixture of term rewriting and code compilation. Our formalization reveals several inconsistencies and contradictions in the current version of the OCL standard. Overall, this document is intended to provide the basis for a machine-checked text "Annex A" of the OCL standard targeting at tool implementors. extra-history = Change history: [2015-10-13]: afp-devel@ea3b38fc54d6 and hol-testgen@12148
   Update of Featherweight OCL including a change in the abstract.
[2014-01-16]: afp-devel@9091ce05cb20 and hol-testgen@10241
   New Entry: Featherweight OCL notify = brucker@spamfence.net, tuong@users.gforge.inria.fr, wolff@lri.fr [Relation_Algebra] title = Relation Algebra author = Alasdair Armstrong <>, Simon Foster , Georg Struth , Tjark Weber date = 2014-01-25 topic = Mathematics/Algebra abstract = Tarski's algebra of binary relations is formalised along the lines of the standard textbooks of Maddux and Schmidt and Ströhlein. This includes relation-algebraic concepts such as subidentities, vectors and a domain operation as well as various notions associated to functions. Relation algebras are also expanded by a reflexive transitive closure operation, and they are linked with Kleene algebras and models of binary relations and Boolean matrices. notify = g.struth@sheffield.ac.uk, tjark.weber@it.uu.se [PSemigroupsConvolution] title = Partial Semigroups and Convolution Algebras author = Brijesh Dongol , Victor B. F. Gomes , Ian J. Hayes , Georg Struth topic = Mathematics/Algebra date = 2017-06-13 notify = g.struth@sheffield.ac.uk, victor.gomes@cl.cam.ac.uk abstract = Partial Semigroups are relevant to the foundations of quantum mechanics and combinatorics as well as to interval and separation logics. Convolution algebras can be understood either as algebras of generalised binary modalities over ternary Kripke frames, in particular over partial semigroups, or as algebras of quantale-valued functions which are equipped with a convolution-style operation of multiplication that is parametrised by a ternary relation. Convolution algebras provide algebraic semantics for various substructural logics, including categorial, relevance and linear logics, for separation logic and for interval logics; they cover quantitative and qualitative applications. These mathematical components for partial semigroups and convolution algebras provide uniform foundations from which models of computation based on relations, program traces or pomsets, and verification components for separation or interval temporal logics can be built with little effort. [Secondary_Sylow] title = Secondary Sylow Theorems author = Jakob von Raumer date = 2014-01-28 topic = Mathematics/Algebra abstract = These theories extend the existing proof of the first Sylow theorem (written by Florian Kammueller and L. C. Paulson) by what are often called the second, third and fourth Sylow theorems. These theorems state propositions about the number of Sylow p-subgroups of a group and the fact that they are conjugate to each other. The proofs make use of an implementation of group actions and their properties. notify = psxjv4@nottingham.ac.uk [Jordan_Hoelder] title = The Jordan-Hölder Theorem author = Jakob von Raumer date = 2014-09-09 topic = Mathematics/Algebra abstract = This submission contains theories that lead to a formalization of the proof of the Jordan-Hölder theorem about composition series of finite groups. The theories formalize the notions of isomorphism classes of groups, simple groups, normal series, composition series, maximal normal subgroups. Furthermore, they provide proofs of the second isomorphism theorem for groups, the characterization theorem for maximal normal subgroups as well as many useful lemmas about normal subgroups and factor groups. The proof is inspired by course notes of Stuart Rankin. notify = psxjv4@nottingham.ac.uk [Cayley_Hamilton] title = The Cayley-Hamilton Theorem author = Stephan Adelsberger , Stefan Hetzl , Florian Pollak date = 2014-09-15 topic = Mathematics/Algebra abstract = This document contains a proof of the Cayley-Hamilton theorem based on the development of matrices in HOL/Multivariate Analysis. notify = stvienna@gmail.com [Probabilistic_Noninterference] title = Probabilistic Noninterference author = Andrei Popescu , Johannes Hölzl date = 2014-03-11 topic = Computer science/Security abstract = We formalize a probabilistic noninterference for a multi-threaded language with uniform scheduling, where probabilistic behaviour comes from both the scheduler and the individual threads. We define notions probabilistic noninterference in two variants: resumption-based and trace-based. For the resumption-based notions, we prove compositionality w.r.t. the language constructs and establish sound type-system-like syntactic criteria. This is a formalization of the mathematical development presented at CPP 2013 and CALCO 2013. It is the probabilistic variant of the Possibilistic Noninterference AFP entry. notify = hoelzl@in.tum.de [HyperCTL] title = A shallow embedding of HyperCTL* author = Markus N. Rabe , Peter Lammich , Andrei Popescu date = 2014-04-16 topic = Computer science/Security, Logic/General logic/Temporal logic abstract = We formalize HyperCTL*, a temporal logic for expressing security properties. We first define a shallow embedding of HyperCTL*, within which we prove inductive and coinductive rules for the operators. Then we show that a HyperCTL* formula captures Goguen-Meseguer noninterference, a landmark information flow property. We also define a deep embedding and connect it to the shallow embedding by a denotational semantics, for which we prove sanity w.r.t. dependence on the free variables. Finally, we show that under some finiteness assumptions about the model, noninterference is given by a (finitary) syntactic formula. notify = uuomul@yahoo.com [Bounded_Deducibility_Security] title = Bounded-Deducibility Security author = Andrei Popescu , Peter Lammich , Thomas Bauereiss date = 2014-04-22 topic = Computer science/Security abstract = This is a formalization of bounded-deducibility security (BD security), a flexible notion of information-flow security applicable to arbitrary transition systems. It generalizes Sutherland's classic notion of nondeducibility by factoring in declassification bounds and trigger, whereas nondeducibility states that, in a system, information cannot flow between specified sources and sinks, BD security indicates upper bounds for the flow and triggers under which these upper bounds are no longer guaranteed. notify = uuomul@yahoo.com, lammich@in.tum.de, thomas@bauereiss.name extra-history = Change history: [2021-08-12]: Generalised BD Security from I/O automata to nondeterministic transition systems, with the former retained as an instance of the latter (renaming locale BD_Security to BD_Security_IO). Generalise unwinding conditions to allow making more than one transition at a time when constructing alternative traces. Add results about the expressivity of declassification triggers vs. bounds, due to Thomas Bauereiss (added as author). [Network_Security_Policy_Verification] title = Network Security Policy Verification author = Cornelius Diekmann date = 2014-07-04 topic = Computer science/Security abstract = We present a unified theory for verifying network security policies. A security policy is represented as directed graph. To check high-level security goals, security invariants over the policy are expressed. We cover monotonic security invariants, i.e. prohibiting more does not harm security. We provide the following contributions for the security invariant theory.
  • Secure auto-completion of scenario-specific knowledge, which eases usability.
  • Security violations can be repaired by tightening the policy iff the security invariants hold for the deny-all policy.
  • An algorithm to compute a security policy.
  • A formalization of stateful connection semantics in network security mechanisms.
  • An algorithm to compute a secure stateful implementation of a policy.
  • An executable implementation of all the theory.
  • Examples, ranging from an aircraft cabin data network to the analysis of a large real-world firewall.
  • More examples: A fully automated translation of high-level security goals to both firewall and SDN configurations (see Examples/Distributed_WebApp.thy).
For a detailed description, see extra-history = Change history: [2015-04-14]: Added Distributed WebApp example and improved graphviz visualization (revision 4dde08ca2ab8)
notify = diekmann@net.in.tum.de [Abstract_Completeness] title = Abstract Completeness author = Jasmin Christian Blanchette , Andrei Popescu , Dmitriy Traytel date = 2014-04-16 topic = Logic/Proof theory abstract = A formalization of an abstract property of possibly infinite derivation trees (modeled by a codatatype), representing the core of a proof (in Beth/Hintikka style) of the first-order logic completeness theorem, independent of the concrete syntax or inference rules. This work is described in detail in the IJCAR 2014 publication by the authors. The abstract proof can be instantiated for a wide range of Gentzen and tableau systems as well as various flavors of FOL---e.g., with or without predicates, equality, or sorts. Here, we give only a toy example instantiation with classical propositional logic. A more serious instance---many-sorted FOL with equality---is described elsewhere [Blanchette and Popescu, FroCoS 2013]. notify = traytel@in.tum.de [Pop_Refinement] title = Pop-Refinement author = Alessandro Coglio date = 2014-07-03 topic = Computer science/Programming languages/Misc abstract = Pop-refinement is an approach to stepwise refinement, carried out inside an interactive theorem prover by constructing a monotonically decreasing sequence of predicates over deeply embedded target programs. The sequence starts with a predicate that characterizes the possible implementations, and ends with a predicate that characterizes a unique program in explicit syntactic form. Pop-refinement enables more requirements (e.g. program-level and non-functional) to be captured in the initial specification and preserved through refinement. Security requirements expressed as hyperproperties (i.e. predicates over sets of traces) are always preserved by pop-refinement, unlike the popular notion of refinement as trace set inclusion. Two simple examples in Isabelle/HOL are presented, featuring program-level requirements, non-functional requirements, and hyperproperties. notify = coglio@kestrel.edu [VectorSpace] title = Vector Spaces author = Holden Lee date = 2014-08-29 topic = Mathematics/Algebra abstract = This formalisation of basic linear algebra is based completely on locales, building off HOL-Algebra. It includes basic definitions: linear combinations, span, linear independence; linear transformations; interpretation of function spaces as vector spaces; the direct sum of vector spaces, sum of subspaces; the replacement theorem; existence of bases in finite-dimensional; vector spaces, definition of dimension; the rank-nullity theorem. Some concepts are actually defined and proved for modules as they also apply there. Infinite-dimensional vector spaces are supported, but dimension is only supported for finite-dimensional vector spaces. The proofs are standard; the proofs of the replacement theorem and rank-nullity theorem roughly follow the presentation in Linear Algebra by Friedberg, Insel, and Spence. The rank-nullity theorem generalises the existing development in the Archive of Formal Proof (originally using type classes, now using a mix of type classes and locales). notify = holdenl@princeton.edu [Special_Function_Bounds] title = Real-Valued Special Functions: Upper and Lower Bounds author = Lawrence C. Paulson date = 2014-08-29 topic = Mathematics/Analysis abstract = This development proves upper and lower bounds for several familiar real-valued functions. For sin, cos, exp and sqrt, it defines and verifies infinite families of upper and lower bounds, mostly based on Taylor series expansions. For arctan, ln and exp, it verifies a finite collection of upper and lower bounds, originally obtained from the functions' continued fraction expansions using the computer algebra system Maple. A common theme in these proofs is to take the difference between a function and its approximation, which should be zero at one point, and then consider the sign of the derivative. The immediate purpose of this development is to verify axioms used by MetiTarski, an automatic theorem prover for real-valued special functions. Crucial to MetiTarski's operation is the provision of upper and lower bounds for each function of interest. notify = lp15@cam.ac.uk [Landau_Symbols] title = Landau Symbols author = Manuel Eberl date = 2015-07-14 topic = Mathematics/Analysis abstract = This entry provides Landau symbols to describe and reason about the asymptotic growth of functions for sufficiently large inputs. A number of simplification procedures are provided for additional convenience: cancelling of dominated terms in sums under a Landau symbol, cancelling of common factors in products, and a decision procedure for Landau expressions containing products of powers of functions like x, ln(x), ln(ln(x)) etc. notify = manuel@pruvisto.org [Error_Function] title = The Error Function author = Manuel Eberl topic = Mathematics/Analysis date = 2018-02-06 notify = manuel@pruvisto.org abstract =

This entry provides the definitions and basic properties of the complex and real error function erf and the complementary error function erfc. Additionally, it gives their full asymptotic expansions.

[Akra_Bazzi] title = The Akra-Bazzi theorem and the Master theorem author = Manuel Eberl date = 2015-07-14 topic = Mathematics/Analysis abstract = This article contains a formalisation of the Akra-Bazzi method based on a proof by Leighton. It is a generalisation of the well-known Master Theorem for analysing the complexity of Divide & Conquer algorithms. We also include a generalised version of the Master theorem based on the Akra-Bazzi theorem, which is easier to apply than the Akra-Bazzi theorem itself.

Some proof methods that facilitate applying the Master theorem are also included. For a more detailed explanation of the formalisation and the proof methods, see the accompanying paper (publication forthcoming). notify = manuel@pruvisto.org [Dirichlet_Series] title = Dirichlet Series author = Manuel Eberl topic = Mathematics/Number theory date = 2017-10-12 notify = manuel@pruvisto.org abstract = This entry is a formalisation of much of Chapters 2, 3, and 11 of Apostol's “Introduction to Analytic Number Theory”. This includes:

  • Definitions and basic properties for several number-theoretic functions (Euler's φ, Möbius μ, Liouville's λ, the divisor function σ, von Mangoldt's Λ)
  • Executable code for most of these functions, the most efficient implementations using the factoring algorithm by Thiemann et al.
  • Dirichlet products and formal Dirichlet series
  • Analytic results connecting convergent formal Dirichlet series to complex functions
  • Euler product expansions
  • Asymptotic estimates of number-theoretic functions including the density of squarefree integers and the average number of divisors of a natural number
These results are useful as a basis for developing more number-theoretic results, such as the Prime Number Theorem. [Gauss_Sums] title = Gauss Sums and the Pólya–Vinogradov Inequality author = Rodrigo Raya , Manuel Eberl topic = Mathematics/Number theory date = 2019-12-10 notify = manuel.eberl@tum.de abstract =

This article provides a full formalisation of Chapter 8 of Apostol's Introduction to Analytic Number Theory. Subjects that are covered are:

  • periodic arithmetic functions and their finite Fourier series
  • (generalised) Ramanujan sums
  • Gauss sums and separable characters
  • induced moduli and primitive characters
  • the Pólya—Vinogradov inequality
[Zeta_Function] title = The Hurwitz and Riemann ζ Functions author = Manuel Eberl topic = Mathematics/Number theory, Mathematics/Analysis date = 2017-10-12 notify = manuel@pruvisto.org abstract =

This entry builds upon the results about formal and analytic Dirichlet series to define the Hurwitz ζ function ζ(a,s) and, based on that, the Riemann ζ function ζ(s). This is done by first defining them for ℜ(z) > 1 and then successively extending the domain to the left using the Euler–MacLaurin formula.

Apart from the most basic facts such as analyticity, the following results are provided:

  • the Stieltjes constants and the Laurent expansion of ζ(s) at s = 1
  • the non-vanishing of ζ(s) for ℜ(z) ≥ 1
  • the relationship between ζ(a,s) and Γ
  • the special values at negative integers and positive even integers
  • Hurwitz's formula and the reflection formula for ζ(s)
  • the Hadjicostas–Chapman formula

The entry also contains Euler's analytic proof of the infinitude of primes, based on the fact that ζ(s) has a pole at s = 1.

[Linear_Recurrences] title = Linear Recurrences author = Manuel Eberl topic = Mathematics/Analysis date = 2017-10-12 notify = manuel@pruvisto.org abstract =

Linear recurrences with constant coefficients are an interesting class of recurrence equations that can be solved explicitly. The most famous example are certainly the Fibonacci numbers with the equation f(n) = f(n-1) + f(n - 2) and the quite non-obvious closed form (φn - (-φ)-n) / √5 where φ is the golden ratio.

In this work, I build on existing tools in Isabelle – such as formal power series and polynomial factorisation algorithms – to develop a theory of these recurrences and derive a fully executable solver for them that can be exported to programming languages like Haskell.

[Van_der_Waerden] title = Van der Waerden's Theorem author = Katharina Kreuzer , Manuel Eberl topic = Mathematics/Combinatorics date = 2021-06-22 notify = kreuzerk@in.tum.de, manuel@pruvisto.org abstract = This article formalises the proof of Van der Waerden's Theorem from Ramsey theory. Van der Waerden's Theorem states that for integers $k$ and $l$ there exists a number $N$ which guarantees that if an integer interval of length at least $N$ is coloured with $k$ colours, there will always be an arithmetic progression of length $l$ of the same colour in said interval. The proof goes along the lines of \cite{Swan}. The smallest number $N_{k,l}$ fulfilling Van der Waerden's Theorem is then called the Van der Waerden Number. Finding the Van der Waerden Number is still an open problem for most values of $k$ and $l$. [Lambert_W] title = The Lambert W Function on the Reals author = Manuel Eberl topic = Mathematics/Analysis date = 2020-04-24 notify = manuel@pruvisto.org abstract =

The Lambert W function is a multi-valued function defined as the inverse function of xx ex. Besides numerous applications in combinatorics, physics, and engineering, it also frequently occurs when solving equations containing both ex and x, or both x and log x.

This article provides a definition of the two real-valued branches W0(x) and W-1(x) and proves various properties such as basic identities and inequalities, monotonicity, differentiability, asymptotic expansions, and the MacLaurin series of W0(x) at x = 0.

[Cartan_FP] title = The Cartan Fixed Point Theorems author = Lawrence C. Paulson date = 2016-03-08 topic = Mathematics/Analysis abstract = The Cartan fixed point theorems concern the group of holomorphic automorphisms on a connected open set of Cn. Ciolli et al. have formalised the one-dimensional case of these theorems in HOL Light. This entry contains their proofs, ported to Isabelle/HOL. Thus it addresses the authors' remark that "it would be important to write a formal proof in a language that can be read by both humans and machines". notify = lp15@cam.ac.uk [Gauss_Jordan] title = Gauss-Jordan Algorithm and Its Applications author = Jose Divasón , Jesús Aransay topic = Computer science/Algorithms/Mathematical date = 2014-09-03 abstract = The Gauss-Jordan algorithm states that any matrix over a field can be transformed by means of elementary row operations to a matrix in reduced row echelon form. The formalization is based on the Rank Nullity Theorem entry of the AFP and on the HOL-Multivariate-Analysis session of Isabelle, where matrices are represented as functions over finite types. We have set up the code generator to make this representation executable. In order to improve the performance, a refinement to immutable arrays has been carried out. We have formalized some of the applications of the Gauss-Jordan algorithm. Thanks to this development, the following facts can be computed over matrices whose elements belong to a field: Ranks, Determinants, Inverses, Bases and dimensions and Solutions of systems of linear equations. Code can be exported to SML and Haskell. notify = jose.divasonm@unirioja.es, jesus-maria.aransay@unirioja.es [Echelon_Form] title = Echelon Form author = Jose Divasón , Jesús Aransay topic = Computer science/Algorithms/Mathematical, Mathematics/Algebra date = 2015-02-12 abstract = We formalize an algorithm to compute the Echelon Form of a matrix. We have proved its existence over Bézout domains and made it executable over Euclidean domains, such as the integer ring and the univariate polynomials over a field. This allows us to compute determinants, inverses and characteristic polynomials of matrices. The work is based on the HOL-Multivariate Analysis library, and on both the Gauss-Jordan and Cayley-Hamilton AFP entries. As a by-product, some algebraic structures have been implemented (principal ideal domains, Bézout domains...). The algorithm has been refined to immutable arrays and code can be generated to functional languages as well. notify = jose.divasonm@unirioja.es, jesus-maria.aransay@unirioja.es [QR_Decomposition] title = QR Decomposition author = Jose Divasón , Jesús Aransay topic = Computer science/Algorithms/Mathematical, Mathematics/Algebra date = 2015-02-12 abstract = QR decomposition is an algorithm to decompose a real matrix A into the product of two other matrices Q and R, where Q is orthogonal and R is invertible and upper triangular. The algorithm is useful for the least squares problem; i.e., the computation of the best approximation of an unsolvable system of linear equations. As a side-product, the Gram-Schmidt process has also been formalized. A refinement using immutable arrays is presented as well. The development relies, among others, on the AFP entry "Implementing field extensions of the form Q[sqrt(b)]" by René Thiemann, which allows execution of the algorithm using symbolic computations. Verified code can be generated and executed using floats as well. extra-history = Change history: [2015-06-18]: The second part of the Fundamental Theorem of Linear Algebra has been generalized to more general inner product spaces. notify = jose.divasonm@unirioja.es, jesus-maria.aransay@unirioja.es [Hermite] title = Hermite Normal Form author = Jose Divasón , Jesús Aransay topic = Computer science/Algorithms/Mathematical, Mathematics/Algebra date = 2015-07-07 abstract = Hermite Normal Form is a canonical matrix analogue of Reduced Echelon Form, but involving matrices over more general rings. In this work we formalise an algorithm to compute the Hermite Normal Form of a matrix by means of elementary row operations, taking advantage of the Echelon Form AFP entry. We have proven the correctness of such an algorithm and refined it to immutable arrays. Furthermore, we have also formalised the uniqueness of the Hermite Normal Form of a matrix. Code can be exported and some examples of execution involving integer matrices and polynomial matrices are presented as well. notify = jose.divasonm@unirioja.es, jesus-maria.aransay@unirioja.es [Imperative_Insertion_Sort] title = Imperative Insertion Sort author = Christian Sternagel date = 2014-09-25 topic = Computer science/Algorithms abstract = The insertion sort algorithm of Cormen et al. (Introduction to Algorithms) is expressed in Imperative HOL and proved to be correct and terminating. For this purpose we also provide a theory about imperative loop constructs with accompanying induction/invariant rules for proving partial and total correctness. Furthermore, the formalized algorithm is fit for code generation. notify = lp15@cam.ac.uk [Stream_Fusion_Code] title = Stream Fusion in HOL with Code Generation author = Andreas Lochbihler , Alexandra Maximova date = 2014-10-10 topic = Computer science/Functional programming abstract = Stream Fusion is a system for removing intermediate list data structures from functional programs, in particular Haskell. This entry adapts stream fusion to Isabelle/HOL and its code generator. We define stream types for finite and possibly infinite lists and stream versions for most of the fusible list functions in the theories List and Coinductive_List, and prove them correct with respect to the conversion functions between lists and streams. The Stream Fusion transformation itself is implemented as a simproc in the preprocessor of the code generator. [Brian Huffman's AFP entry formalises stream fusion in HOLCF for the domain of lazy lists to prove the GHC compiler rewrite rules correct. In contrast, this work enables Isabelle's code generator to perform stream fusion itself. To that end, it covers both finite and coinductive lists from the HOL library and the Coinductive entry. The fusible list functions require specification and proof principles different from Huffman's.] notify = mail@andreas-lochbihler.de [Case_Labeling] title = Generating Cases from Labeled Subgoals author = Lars Noschinski date = 2015-07-21 topic = Tools, Computer science/Programming languages/Misc abstract = Isabelle/Isar provides named cases to structure proofs. This article contains an implementation of a proof method casify, which can be used to easily extend proof tools with support for named cases. Such a proof tool must produce labeled subgoals, which are then interpreted by casify.

As examples, this work contains verification condition generators producing named cases for three languages: The Hoare language from HOL/Library, a monadic language for computations with failure (inspired by the AutoCorres tool), and a language of conditional expressions. These VCGs are demonstrated by a number of example programs. notify = noschinl@gmail.com [DPT-SAT-Solver] title = A Fast SAT Solver for Isabelle in Standard ML topic = Tools author = Armin Heller <> date = 2009-12-09 abstract = This contribution contains a fast SAT solver for Isabelle written in Standard ML. By loading the theory DPT_SAT_Solver, the SAT solver installs itself (under the name ``dptsat'') and certain Isabelle tools like Refute will start using it automatically. This is a port of the DPT (Decision Procedure Toolkit) SAT Solver written in OCaml. notify = jasmin.blanchette@gmail.com [Rep_Fin_Groups] title = Representations of Finite Groups topic = Mathematics/Algebra author = Jeremy Sylvestre date = 2015-08-12 abstract = We provide a formal framework for the theory of representations of finite groups, as modules over the group ring. Along the way, we develop the general theory of groups (relying on the group_add class for the basics), modules, and vector spaces, to the extent required for theory of group representations. We then provide formal proofs of several important introductory theorems in the subject, including Maschke's theorem, Schur's lemma, and Frobenius reciprocity. We also prove that every irreducible representation is isomorphic to a submodule of the group ring, leading to the fact that for a finite group there are only finitely many isomorphism classes of irreducible representations. In all of this, no restriction is made on the characteristic of the ring or field of scalars until the definition of a group representation, and then the only restriction made is that the characteristic must not divide the order of the group. notify = jsylvest@ualberta.ca [Noninterference_Inductive_Unwinding] title = The Inductive Unwinding Theorem for CSP Noninterference Security topic = Computer science/Security author = Pasquale Noce date = 2015-08-18 abstract =

The necessary and sufficient condition for CSP noninterference security stated by the Ipurge Unwinding Theorem is expressed in terms of a pair of event lists varying over the set of process traces. This does not render it suitable for the subsequent application of rule induction in the case of a process defined inductively, since rule induction may rather be applied to a single variable ranging over an inductively defined set.

Starting from the Ipurge Unwinding Theorem, this paper derives a necessary and sufficient condition for CSP noninterference security that involves a single event list varying over the set of process traces, and is thus suitable for rule induction; hence its name, Inductive Unwinding Theorem. Similarly to the Ipurge Unwinding Theorem, the new theorem only requires to consider individual accepted and refused events for each process trace, and applies to the general case of a possibly intransitive noninterference policy. Specific variants of this theorem are additionally proven for deterministic processes and trace set processes.

notify = pasquale.noce.lavoro@gmail.com [Password_Authentication_Protocol] title = Verification of a Diffie-Hellman Password-based Authentication Protocol by Extending the Inductive Method author = Pasquale Noce topic = Computer science/Security date = 2017-01-03 notify = pasquale.noce.lavoro@gmail.com abstract = This paper constructs a formal model of a Diffie-Hellman password-based authentication protocol between a user and a smart card, and proves its security. The protocol provides for the dispatch of the user's password to the smart card on a secure messaging channel established by means of Password Authenticated Connection Establishment (PACE), where the mapping method being used is Chip Authentication Mapping. By applying and suitably extending Paulson's Inductive Method, this paper proves that the protocol establishes trustworthy secure messaging channels, preserves the secrecy of users' passwords, and provides an effective mutual authentication service. What is more, these security properties turn out to hold independently of the secrecy of the PACE authentication key. [Jordan_Normal_Form] title = Matrices, Jordan Normal Forms, and Spectral Radius Theory topic = Mathematics/Algebra author = René Thiemann , Akihisa Yamada contributors = Alexander Bentkamp date = 2015-08-21 abstract =

Matrix interpretations are useful as measure functions in termination proving. In order to use these interpretations also for complexity analysis, the growth rate of matrix powers has to examined. Here, we formalized a central result of spectral radius theory, namely that the growth rate is polynomially bounded if and only if the spectral radius of a matrix is at most one.

To formally prove this result we first studied the growth rates of matrices in Jordan normal form, and prove the result that every complex matrix has a Jordan normal form using a constructive prove via Schur decomposition.

The whole development is based on a new abstract type for matrices, which is also executable by a suitable setup of the code generator. It completely subsumes our former AFP-entry on executable matrices, and its main advantage is its close connection to the HMA-representation which allowed us to easily adapt existing proofs on determinants.

All the results have been applied to improve CeTA, our certifier to validate termination and complexity proof certificates.

extra-history = Change history: [2016-01-07]: Added Schur-decomposition, Gram-Schmidt orthogonalization, uniqueness of Jordan normal forms
[2018-04-17]: Integrated lemmas from deep-learning AFP-entry of Alexander Bentkamp notify = rene.thiemann@uibk.ac.at, ayamada@trs.cm.is.nagoya-u.ac.jp [LTL_to_DRA] title = Converting Linear Temporal Logic to Deterministic (Generalized) Rabin Automata topic = Computer science/Automata and formal languages author = Salomon Sickert date = 2015-09-04 abstract = Recently, Javier Esparza and Jan Kretinsky proposed a new method directly translating linear temporal logic (LTL) formulas to deterministic (generalized) Rabin automata. Compared to the existing approaches of constructing a non-deterministic Buechi-automaton in the first step and then applying a determinization procedure (e.g. some variant of Safra's construction) in a second step, this new approach preservers a relation between the formula and the states of the resulting automaton. While the old approach produced a monolithic structure, the new method is compositional. Furthermore, in some cases the resulting automata are much smaller than the automata generated by existing approaches. In order to ensure the correctness of the construction, this entry contains a complete formalisation and verification of the translation. Furthermore from this basis executable code is generated. extra-history = Change history: [2015-09-23]: Enable code export for the eager unfolding optimisation and reduce running time of the generated tool. Moreover, add support for the mlton SML compiler.
[2016-03-24]: Make use of the LTL entry and include the simplifier. notify = sickert@in.tum.de [Timed_Automata] title = Timed Automata author = Simon Wimmer date = 2016-03-08 topic = Computer science/Automata and formal languages abstract = Timed automata are a widely used formalism for modeling real-time systems, which is employed in a class of successful model checkers such as UPPAAL [LPY97], HyTech [HHWt97] or Kronos [Yov97]. This work formalizes the theory for the subclass of diagonal-free timed automata, which is sufficient to model many interesting problems. We first define the basic concepts and semantics of diagonal-free timed automata. Based on this, we prove two types of decidability results for the language emptiness problem. The first is the classic result of Alur and Dill [AD90, AD94], which uses a finite partitioning of the state space into so-called `regions`. Our second result focuses on an approach based on `Difference Bound Matrices (DBMs)`, which is practically used by model checkers. We prove the correctness of the basic forward analysis operations on DBMs. One of these operations is the Floyd-Warshall algorithm for the all-pairs shortest paths problem. To obtain a finite search space, a widening operation has to be used for this kind of analysis. We use Patricia Bouyer's [Bou04] approach to prove that this widening operation is correct in the sense that DBM-based forward analysis in combination with the widening operation also decides language emptiness. The interesting property of this proof is that the first decidability result is reused to obtain the second one. notify = wimmers@in.tum.de [Parity_Game] title = Positional Determinacy of Parity Games author = Christoph Dittmann date = 2015-11-02 topic = Mathematics/Games and economics, Mathematics/Graph theory abstract = We present a formalization of parity games (a two-player game on directed graphs) and a proof of their positional determinacy in Isabelle/HOL. This proof works for both finite and infinite games. notify = [Ergodic_Theory] title = Ergodic Theory author = Sebastien Gouezel contributors = Manuel Eberl date = 2015-12-01 topic = Mathematics/Probability theory abstract = Ergodic theory is the branch of mathematics that studies the behaviour of measure preserving transformations, in finite or infinite measure. It interacts both with probability theory (mainly through measure theory) and with geometry as a lot of interesting examples are from geometric origin. We implement the first definitions and theorems of ergodic theory, including notably Poicaré recurrence theorem for finite measure preserving systems (together with the notion of conservativity in general), induced maps, Kac's theorem, Birkhoff theorem (arguably the most important theorem in ergodic theory), and variations around it such as conservativity of the corresponding skew product, or Atkinson lemma. notify = sebastien.gouezel@univ-rennes1.fr, hoelzl@in.tum.de [Latin_Square] title = Latin Square author = Alexander Bentkamp date = 2015-12-02 topic = Mathematics/Combinatorics abstract = A Latin Square is a n x n table filled with integers from 1 to n where each number appears exactly once in each row and each column. A Latin Rectangle is a partially filled n x n table with r filled rows and n-r empty rows, such that each number appears at most once in each row and each column. The main result of this theory is that any Latin Rectangle can be completed to a Latin Square. notify = bentkamp@gmail.com [Deep_Learning] title = Expressiveness of Deep Learning author = Alexander Bentkamp date = 2016-11-10 topic = Computer science/Machine learning, Mathematics/Analysis abstract = Deep learning has had a profound impact on computer science in recent years, with applications to search engines, image recognition and language processing, bioinformatics, and more. Recently, Cohen et al. provided theoretical evidence for the superiority of deep learning over shallow learning. This formalization of their work simplifies and generalizes the original proof, while working around the limitations of the Isabelle type system. To support the formalization, I developed reusable libraries of formalized mathematics, including results about the matrix rank, the Lebesgue measure, and multivariate polynomials, as well as a library for tensor analysis. notify = bentkamp@gmail.com [Inductive_Inference] title = Some classical results in inductive inference of recursive functions author = Frank J. Balbach topic = Logic/Computability, Computer science/Machine learning date = 2020-08-31 notify = frank-balbach@gmx.de abstract =

This entry formalizes some classical concepts and results from inductive inference of recursive functions. In the basic setting a partial recursive function ("strategy") must identify ("learn") all functions from a set ("class") of recursive functions. To that end the strategy receives more and more values $f(0), f(1), f(2), \ldots$ of some function $f$ from the given class and in turn outputs descriptions of partial recursive functions, for example, Gödel numbers. The strategy is considered successful if the sequence of outputs ("hypotheses") converges to a description of $f$. A class of functions learnable in this sense is called "learnable in the limit". The set of all these classes is denoted by LIM.

Other types of inference considered are finite learning (FIN), behaviorally correct learning in the limit (BC), and some variants of LIM with restrictions on the hypotheses: total learning (TOTAL), consistent learning (CONS), and class-preserving learning (CP). The main results formalized are the proper inclusions $\mathrm{FIN} \subset \mathrm{CP} \subset \mathrm{TOTAL} \subset \mathrm{CONS} \subset \mathrm{LIM} \subset \mathrm{BC} \subset 2^{\mathcal{R}}$, where $\mathcal{R}$ is the set of all total recursive functions. Further results show that for all these inference types except CONS, strategies can be assumed to be total recursive functions; that all inference types but CP are closed under the subset relation between classes; and that no inference type is closed under the union of classes.

The above is based on a formalization of recursive functions heavily inspired by the Universal Turing Machine entry by Xu et al., but different in that it models partial functions with codomain nat option. The formalization contains a construction of a universal partial recursive function, without resorting to Turing machines, introduces decidability and recursive enumerability, and proves some standard results: existence of a Kleene normal form, the s-m-n theorem, Rice's theorem, and assorted fixed-point theorems (recursion theorems) by Kleene, Rogers, and Smullyan.

[Applicative_Lifting] title = Applicative Lifting author = Andreas Lochbihler , Joshua Schneider <> date = 2015-12-22 topic = Computer science/Functional programming abstract = Applicative functors augment computations with effects by lifting function application to types which model the effects. As the structure of the computation cannot depend on the effects, applicative expressions can be analysed statically. This allows us to lift universally quantified equations to the effectful types, as observed by Hinze. Thus, equational reasoning over effectful computations can be reduced to pure types.

This entry provides a package for registering applicative functors and two proof methods for lifting of equations over applicative functors. The first method normalises applicative expressions according to the laws of applicative functors. This way, equations whose two sides contain the same list of variables can be lifted to every applicative functor.

To lift larger classes of equations, the second method exploits a number of additional properties (e.g., commutativity of effects) provided the properties have been declared for the concrete applicative functor at hand upon registration.

We declare several types from the Isabelle library as applicative functors and illustrate the use of the methods with two examples: the lifting of the arithmetic type class hierarchy to streams and the verification of a relabelling function on binary trees. We also formalise and verify the normalisation algorithm used by the first proof method.

extra-history = Change history: [2016-03-03]: added formalisation of lifting with combinators
[2016-06-10]: implemented automatic derivation of lifted combinator reductions; support arbitrary lifted relations using relators; improved compatibility with locale interpretation (revision ec336f354f37)
notify = mail@andreas-lochbihler.de [Stern_Brocot] title = The Stern-Brocot Tree author = Peter Gammie , Andreas Lochbihler date = 2015-12-22 topic = Mathematics/Number theory abstract = The Stern-Brocot tree contains all rational numbers exactly once and in their lowest terms. We formalise the Stern-Brocot tree as a coinductive tree using recursive and iterative specifications, which we have proven equivalent, and show that it indeed contains all the numbers as stated. Following Hinze, we prove that the Stern-Brocot tree can be linearised looplessly into Stern's diatonic sequence (also known as Dijkstra's fusc function) and that it is a permutation of the Bird tree.

The reasoning stays at an abstract level by appealing to the uniqueness of solutions of guarded recursive equations and lifting algebraic laws point-wise to trees and streams using applicative functors.

notify = mail@andreas-lochbihler.de [Algebraic_Numbers] title = Algebraic Numbers in Isabelle/HOL topic = Mathematics/Algebra author = René Thiemann , Akihisa Yamada , Sebastiaan Joosten contributors = Manuel Eberl date = 2015-12-22 abstract = Based on existing libraries for matrices, factorization of rational polynomials, and Sturm's theorem, we formalized algebraic numbers in Isabelle/HOL. Our development serves as an implementation for real and complex numbers, and it admits to compute roots and completely factorize real and complex polynomials, provided that all coefficients are rational numbers. Moreover, we provide two implementations to display algebraic numbers, an injective and expensive one, or a faster but approximative version.

To this end, we mechanized several results on resultants, which also required us to prove that polynomials over a unique factorization domain form again a unique factorization domain.

extra-history = Change history: [2016-01-29]: Split off Polynomial Interpolation and Polynomial Factorization
[2017-04-16]: Use certified Berlekamp-Zassenhaus factorization, use subresultant algorithm for computing resultants, improved bisection algorithm notify = rene.thiemann@uibk.ac.at, ayamada@trs.cm.is.nagoya-u.ac.jp, sebastiaan.joosten@uibk.ac.at [Polynomial_Interpolation] title = Polynomial Interpolation topic = Mathematics/Algebra author = René Thiemann , Akihisa Yamada date = 2016-01-29 abstract = We formalized three algorithms for polynomial interpolation over arbitrary fields: Lagrange's explicit expression, the recursive algorithm of Neville and Aitken, and the Newton interpolation in combination with an efficient implementation of divided differences. Variants of these algorithms for integer polynomials are also available, where sometimes the interpolation can fail; e.g., there is no linear integer polynomial p such that p(0) = 0 and p(2) = 1. Moreover, for the Newton interpolation for integer polynomials, we proved that all intermediate results that are computed during the algorithm must be integers. This admits an early failure detection in the implementation. Finally, we proved the uniqueness of polynomial interpolation.

The development also contains improved code equations to speed up the division of integers in target languages. notify = rene.thiemann@uibk.ac.at, ayamada@trs.cm.is.nagoya-u.ac.jp [Polynomial_Factorization] title = Polynomial Factorization topic = Mathematics/Algebra author = René Thiemann , Akihisa Yamada date = 2016-01-29 abstract = Based on existing libraries for polynomial interpolation and matrices, we formalized several factorization algorithms for polynomials, including Kronecker's algorithm for integer polynomials, Yun's square-free factorization algorithm for field polynomials, and Berlekamp's algorithm for polynomials over finite fields. By combining the last one with Hensel's lifting, we derive an efficient factorization algorithm for the integer polynomials, which is then lifted for rational polynomials by mechanizing Gauss' lemma. Finally, we assembled a combined factorization algorithm for rational polynomials, which combines all the mentioned algorithms and additionally uses the explicit formula for roots of quadratic polynomials and a rational root test.

As side products, we developed division algorithms for polynomials over integral domains, as well as primality-testing and prime-factorization algorithms for integers. notify = rene.thiemann@uibk.ac.at, ayamada@trs.cm.is.nagoya-u.ac.jp [Cubic_Quartic_Equations] title = Solving Cubic and Quartic Equations author = René Thiemann topic = Mathematics/Analysis date = 2021-09-03 notify = rene.thiemann@uibk.ac.at abstract =

We formalize Cardano's formula to solve a cubic equation $$ax^3 + bx^2 + cx + d = 0,$$ as well as Ferrari's formula to solve a quartic equation. We further turn both formulas into executable algorithms based on the algebraic number implementation in the AFP. To this end we also slightly extended this library, namely by making the minimal polynomial of an algebraic number executable, and by defining and implementing $n$-th roots of complex numbers.

[Perron_Frobenius] title = Perron-Frobenius Theorem for Spectral Radius Analysis author = Jose Divasón , Ondřej Kunčar , René Thiemann , Akihisa Yamada notify = rene.thiemann@uibk.ac.at date = 2016-05-20 topic = Mathematics/Algebra abstract =

The spectral radius of a matrix A is the maximum norm of all eigenvalues of A. In previous work we already formalized that for a complex matrix A, the values in An grow polynomially in n if and only if the spectral radius is at most one. One problem with the above characterization is the determination of all complex eigenvalues. In case A contains only non-negative real values, a simplification is possible with the help of the Perron–Frobenius theorem, which tells us that it suffices to consider only the real eigenvalues of A, i.e., applying Sturm's method can decide the polynomial growth of An.

We formalize the Perron–Frobenius theorem based on a proof via Brouwer's fixpoint theorem, which is available in the HOL multivariate analysis (HMA) library. Since the results on the spectral radius is based on matrices in the Jordan normal form (JNF) library, we further develop a connection which allows us to easily transfer theorems between HMA and JNF. With this connection we derive the combined result: if A is a non-negative real matrix, and no real eigenvalue of A is strictly larger than one, then An is polynomially bounded in n.

extra-history = Change history: [2017-10-18]: added Perron-Frobenius theorem for irreducible matrices with generalization (revision bda1f1ce8a1c)
[2018-05-17]: prove conjecture of CPP'18 paper: Jordan blocks of spectral radius have maximum size (revision ffdb3794e5d5) [Stochastic_Matrices] title = Stochastic Matrices and the Perron-Frobenius Theorem author = René Thiemann topic = Mathematics/Algebra, Computer science/Automata and formal languages date = 2017-11-22 notify = rene.thiemann@uibk.ac.at abstract = Stochastic matrices are a convenient way to model discrete-time and finite state Markov chains. The Perron–Frobenius theorem tells us something about the existence and uniqueness of non-negative eigenvectors of a stochastic matrix. In this entry, we formalize stochastic matrices, link the formalization to the existing AFP-entry on Markov chains, and apply the Perron–Frobenius theorem to prove that stationary distributions always exist, and they are unique if the stochastic matrix is irreducible. [Formal_SSA] title = Verified Construction of Static Single Assignment Form author = Sebastian Ullrich , Denis Lohner date = 2016-02-05 topic = Computer science/Algorithms, Computer science/Programming languages/Transformations abstract =

We define a functional variant of the static single assignment (SSA) form construction algorithm described by Braun et al., which combines simplicity and efficiency. The definition is based on a general, abstract control flow graph representation using Isabelle locales.

We prove that the algorithm's output is semantically equivalent to the input according to a small-step semantics, and that it is in minimal SSA form for the common special case of reducible inputs. We then show the satisfiability of the locale assumptions by giving instantiations for a simple While language.

Furthermore, we use a generic instantiation based on typedefs in order to extract OCaml code and replace the unverified SSA construction algorithm of the CompCertSSA project with it.

A more detailed description of the verified SSA construction can be found in the paper Verified Construction of Static Single Assignment Form, CC 2016.

notify = denis.lohner@kit.edu [Minimal_SSA] title = Minimal Static Single Assignment Form author = Max Wagner , Denis Lohner topic = Computer science/Programming languages/Transformations date = 2017-01-17 notify = denis.lohner@kit.edu abstract =

This formalization is an extension to "Verified Construction of Static Single Assignment Form". In their work, the authors have shown that Braun et al.'s static single assignment (SSA) construction algorithm produces minimal SSA form for input programs with a reducible control flow graph (CFG). However Braun et al. also proposed an extension to their algorithm that they claim produces minimal SSA form even for irreducible CFGs.
In this formalization we support that claim by giving a mechanized proof.

As the extension of Braun et al.'s algorithm aims for removing so-called redundant strongly connected components of phi functions, we show that this suffices to guarantee minimality according to Cytron et al..

[PropResPI] title = Propositional Resolution and Prime Implicates Generation author = Nicolas Peltier notify = Nicolas.Peltier@imag.fr date = 2016-03-11 topic = Logic/General logic/Mechanization of proofs abstract = We provide formal proofs in Isabelle-HOL (using mostly structured Isar proofs) of the soundness and completeness of the Resolution rule in propositional logic. The completeness proofs take into account the usual redundancy elimination rules (tautology elimination and subsumption), and several refinements of the Resolution rule are considered: ordered resolution (with selection functions), positive and negative resolution, semantic resolution and unit resolution (the latter refinement is complete only for clause sets that are Horn- renamable). We also define a concrete procedure for computing saturated sets and establish its soundness and completeness. The clause sets are not assumed to be finite, so that the results can be applied to formulas obtained by grounding sets of first-order clauses (however, a total ordering among atoms is assumed to be given). Next, we show that the unrestricted Resolution rule is deductive- complete, in the sense that it is able to generate all (prime) implicates of any set of propositional clauses (i.e., all entailment- minimal, non-valid, clausal consequences of the considered set). The generation of prime implicates is an important problem, with many applications in artificial intelligence and verification (for abductive reasoning, knowledge compilation, diagnosis, debugging etc.). We also show that implicates can be computed in an incremental way, by fixing an ordering among all the atoms in the considered sets and resolving upon these atoms one by one in the considered order (with no backtracking). This feature is critical for the efficient computation of prime implicates. Building on these results, we provide a procedure for computing such implicates and establish its soundness and completeness. [SuperCalc] title = A Variant of the Superposition Calculus author = Nicolas Peltier notify = Nicolas.Peltier@imag.fr date = 2016-09-06 topic = Logic/Proof theory abstract = We provide a formalization of a variant of the superposition calculus, together with formal proofs of soundness and refutational completeness (w.r.t. the usual redundancy criteria based on clause ordering). This version of the calculus uses all the standard restrictions of the superposition rules, together with the following refinement, inspired by the basic superposition calculus: each clause is associated with a set of terms which are assumed to be in normal form -- thus any application of the replacement rule on these terms is blocked. The set is initially empty and terms may be added or removed at each inference step. The set of terms that are assumed to be in normal form includes any term introduced by previous unifiers as well as any term occurring in the parent clauses at a position that is smaller (according to some given ordering on positions) than a previously replaced term. The standard superposition calculus corresponds to the case where the set of irreducible terms is always empty. [Nominal2] title = Nominal 2 author = Christian Urban , Stefan Berghofer , Cezary Kaliszyk date = 2013-02-21 topic = Tools abstract =

Dealing with binders, renaming of bound variables, capture-avoiding substitution, etc., is very often a major problem in formal proofs, especially in proofs by structural and rule induction. Nominal Isabelle is designed to make such proofs easy to formalise: it provides an infrastructure for declaring nominal datatypes (that is alpha-equivalence classes) and for defining functions over them by structural recursion. It also provides induction principles that have Barendregt’s variable convention already built in.

This entry can be used as a more advanced replacement for HOL/Nominal in the Isabelle distribution.

notify = christian.urban@kcl.ac.uk [First_Welfare_Theorem] title = Microeconomics and the First Welfare Theorem author = Julian Parsert , Cezary Kaliszyk topic = Mathematics/Games and economics license = LGPL date = 2017-09-01 notify = julian.parsert@uibk.ac.at, cezary.kaliszyk@uibk.ac.at abstract = Economic activity has always been a fundamental part of society. Due to modern day politics, economic theory has gained even more influence on our lives. Thus we want models and theories to be as precise as possible. This can be achieved using certification with the help of formal proof technology. Hence we will use Isabelle/HOL to construct two economic models, that of the the pure exchange economy and a version of the Arrow-Debreu Model. We will prove that the First Theorem of Welfare Economics holds within both. The theorem is the mathematical formulation of Adam Smith's famous invisible hand and states that a group of self-interested and rational actors will eventually achieve an efficient allocation of goods and services. extra-history = Change history: [2018-06-17]: Added some lemmas and a theory file, also introduced Microeconomics folder.
[Noninterference_Sequential_Composition] title = Conservation of CSP Noninterference Security under Sequential Composition author = Pasquale Noce date = 2016-04-26 topic = Computer science/Security, Computer science/Concurrency/Process calculi abstract =

In his outstanding work on Communicating Sequential Processes, Hoare has defined two fundamental binary operations allowing to compose the input processes into another, typically more complex, process: sequential composition and concurrent composition. Particularly, the output of the former operation is a process that initially behaves like the first operand, and then like the second operand once the execution of the first one has terminated successfully, as long as it does.

This paper formalizes Hoare's definition of sequential composition and proves, in the general case of a possibly intransitive policy, that CSP noninterference security is conserved under this operation, provided that successful termination cannot be affected by confidential events and cannot occur as an alternative to other events in the traces of the first operand. Both of these assumptions are shown, by means of counterexamples, to be necessary for the theorem to hold.

notify = pasquale.noce.lavoro@gmail.com [Noninterference_Concurrent_Composition] title = Conservation of CSP Noninterference Security under Concurrent Composition author = Pasquale Noce notify = pasquale.noce.lavoro@gmail.com date = 2016-06-13 topic = Computer science/Security, Computer science/Concurrency/Process calculi abstract =

In his outstanding work on Communicating Sequential Processes, Hoare has defined two fundamental binary operations allowing to compose the input processes into another, typically more complex, process: sequential composition and concurrent composition. Particularly, the output of the latter operation is a process in which any event not shared by both operands can occur whenever the operand that admits the event can engage in it, whereas any event shared by both operands can occur just in case both can engage in it.

This paper formalizes Hoare's definition of concurrent composition and proves, in the general case of a possibly intransitive policy, that CSP noninterference security is conserved under this operation. This result, along with the previous analogous one concerning sequential composition, enables the construction of more and more complex processes enforcing noninterference security by composing, sequentially or concurrently, simpler secure processes, whose security can in turn be proven using either the definition of security, or unwinding theorems.

[ROBDD] title = Algorithms for Reduced Ordered Binary Decision Diagrams author = Julius Michaelis , Maximilian Haslbeck , Peter Lammich , Lars Hupel date = 2016-04-27 topic = Computer science/Algorithms, Computer science/Data structures abstract = We present a verified and executable implementation of ROBDDs in Isabelle/HOL. Our implementation relates pointer-based computation in the Heap monad to operations on an abstract definition of boolean functions. Internally, we implemented the if-then-else combinator in a recursive fashion, following the Shannon decomposition of the argument functions. The implementation mixes and adapts known techniques and is built with efficiency in mind. notify = bdd@liftm.de, haslbecm@in.tum.de [No_FTL_observers] title = No Faster-Than-Light Observers author = Mike Stannett , István Németi date = 2016-04-28 topic = Mathematics/Physics abstract = We provide a formal proof within First Order Relativity Theory that no observer can travel faster than the speed of light. Originally reported in Stannett & Németi (2014) "Using Isabelle/HOL to verify first-order relativity theory", Journal of Automated Reasoning 52(4), pp. 361-378. notify = m.stannett@sheffield.ac.uk [Schutz_Spacetime] title = Schutz' Independent Axioms for Minkowski Spacetime -author = Richard Schmoetten , Jake Palmer , Jacques Fleuriot +author = Richard Schmoetten , Jake Palmer , Jacques Fleuriot topic = Mathematics/Physics, Mathematics/Geometry date = 2021-07-27 notify = s1311325@sms.ed.ac.uk abstract = This is a formalisation of Schutz' system of axioms for Minkowski spacetime published under the name "Independent axioms for Minkowski space-time" in 1997, as well as most of the results in the third chapter ("Temporal Order on a Path") of the above monograph. Many results are proven here that cannot be found in Schutz, either preceding the theorem they are needed for, or within their own thematic section. +[Real_Power] +title = Real Exponents as the Limits of Sequences of Rational Exponents +author = Jacques D. Fleuriot +topic = Mathematics/Analysis +date = 2021-11-08 +notify = jdf@ed.ac.uk +abstract = + In this formalisation, we construct real exponents as the limits of + sequences of rational exponents. In particular, if $a \ge 1$ and $x + \in \mathbb{R}$, we choose an increasing rational sequence $r_n$ such + that $\lim_{n\to\infty} {r_n} = x$. Then the sequence $a^{r_n}$ is + increasing and if $r$ is any rational number such that $r > x$, + $a^{r_n}$ is bounded above by $a^r$. By the convergence criterion for + monotone sequences, $a^{r_n}$ converges. We define $a^ x = + \lim_{n\to\infty} a^{r_n}$ and show that it has the expected + properties (for $a \ge 0$). This particular construction of real + exponents is needed instead of the usual one using the natural + logarithm and exponential functions (which already exists in Isabelle) + to support our mechanical derivation of Euler's exponential + series as an ``infinite polynomial". Aside from helping us avoid + circular reasoning, this is, as far as we are aware, the first time + real exponents are mechanised in this way within a proof assistant. + [Groebner_Bases] title = Gröbner Bases Theory author = Fabian Immler , Alexander Maletzky date = 2016-05-02 topic = Mathematics/Algebra, Computer science/Algorithms/Mathematical abstract = This formalization is concerned with the theory of Gröbner bases in (commutative) multivariate polynomial rings over fields, originally developed by Buchberger in his 1965 PhD thesis. Apart from the statement and proof of the main theorem of the theory, the formalization also implements Buchberger's algorithm for actually computing Gröbner bases as a tail-recursive function, thus allowing to effectively decide ideal membership in finitely generated polynomial ideals. Furthermore, all functions can be executed on a concrete representation of multivariate polynomials as association lists. extra-history = Change history: [2019-04-18]: Specialized Gröbner bases to less abstract representation of polynomials, where power-products are represented as polynomial mappings.
notify = alexander.maletzky@risc.jku.at [Nullstellensatz] title = Hilbert's Nullstellensatz author = Alexander Maletzky topic = Mathematics/Algebra, Mathematics/Geometry date = 2019-06-16 notify = alexander.maletzky@risc-software.at abstract = This entry formalizes Hilbert's Nullstellensatz, an important theorem in algebraic geometry that can be viewed as the generalization of the Fundamental Theorem of Algebra to multivariate polynomials: If a set of (multivariate) polynomials over an algebraically closed field has no common zero, then the ideal it generates is the entire polynomial ring. The formalization proves several equivalent versions of this celebrated theorem: the weak Nullstellensatz, the strong Nullstellensatz (connecting algebraic varieties and radical ideals), and the field-theoretic Nullstellensatz. The formalization follows Chapter 4.1. of Ideals, Varieties, and Algorithms by Cox, Little and O'Shea. [Bell_Numbers_Spivey] title = Spivey's Generalized Recurrence for Bell Numbers author = Lukas Bulwahn date = 2016-05-04 topic = Mathematics/Combinatorics abstract = This entry defines the Bell numbers as the cardinality of set partitions for a carrier set of given size, and derives Spivey's generalized recurrence relation for Bell numbers following his elegant and intuitive combinatorial proof.

As the set construction for the combinatorial proof requires construction of three intermediate structures, the main difficulty of the formalization is handling the overall combinatorial argument in a structured way. The introduced proof structure allows us to compose the combinatorial argument from its subparts, and supports to keep track how the detailed proof steps are related to the overall argument. To obtain this structure, this entry uses set monad notation for the set construction's definition, introduces suitable predicates and rules, and follows a repeating structure in its Isar proof. notify = lukas.bulwahn@gmail.com [Randomised_Social_Choice] title = Randomised Social Choice Theory author = Manuel Eberl date = 2016-05-05 topic = Mathematics/Games and economics abstract = This work contains a formalisation of basic Randomised Social Choice, including Stochastic Dominance and Social Decision Schemes (SDSs) along with some of their most important properties (Anonymity, Neutrality, ex-post- and SD-Efficiency, SD-Strategy-Proofness) and two particular SDSs – Random Dictatorship and Random Serial Dictatorship (with proofs of the properties that they satisfy). Many important properties of these concepts are also proven – such as the two equivalent characterisations of Stochastic Dominance and the fact that SD-efficiency of a lottery only depends on the support. The entry also provides convenient commands to define Preference Profiles, prove their well-formedness, and automatically derive restrictions that sufficiently nice SDSs need to satisfy on the defined profiles. Currently, the formalisation focuses on weak preferences and Stochastic Dominance, but it should be easy to extend it to other domains – such as strict preferences – or other lottery extensions – such as Bilinear Dominance or Pairwise Comparison. notify = manuel@pruvisto.org [SDS_Impossibility] title = The Incompatibility of SD-Efficiency and SD-Strategy-Proofness author = Manuel Eberl date = 2016-05-04 topic = Mathematics/Games and economics abstract = This formalisation contains the proof that there is no anonymous and neutral Social Decision Scheme for at least four voters and alternatives that fulfils both SD-Efficiency and SD-Strategy- Proofness. The proof is a fully structured and quasi-human-redable one. It was derived from the (unstructured) SMT proof of the case for exactly four voters and alternatives by Brandl et al. Their proof relies on an unverified translation of the original problem to SMT, and the proof that lifts the argument for exactly four voters and alternatives to the general case is also not machine-checked. In this Isabelle proof, on the other hand, all of these steps are fully proven and machine-checked. This is particularly important seeing as a previously published informal proof of a weaker statement contained a mistake in precisely this lifting step. notify = manuel@pruvisto.org [Median_Of_Medians_Selection] title = The Median-of-Medians Selection Algorithm author = Manuel Eberl topic = Computer science/Algorithms date = 2017-12-21 notify = manuel@pruvisto.org abstract =

This entry provides an executable functional implementation of the Median-of-Medians algorithm for selecting the k-th smallest element of an unsorted list deterministically in linear time. The size bounds for the recursive call that lead to the linear upper bound on the run-time of the algorithm are also proven.

[Mason_Stothers] title = The Mason–Stothers Theorem author = Manuel Eberl topic = Mathematics/Algebra date = 2017-12-21 notify = manuel@pruvisto.org abstract =

This article provides a formalisation of Snyder’s simple and elegant proof of the Mason–Stothers theorem, which is the polynomial analogue of the famous abc Conjecture for integers. Remarkably, Snyder found this very elegant proof when he was still a high-school student.

In short, the statement of the theorem is that three non-zero coprime polynomials A, B, C over a field which sum to 0 and do not all have vanishing derivatives fulfil max{deg(A), deg(B), deg(C)} < deg(rad(ABC)) where the rad(P) denotes the radical of P, i. e. the product of all unique irreducible factors of P.

This theorem also implies a kind of polynomial analogue of Fermat’s Last Theorem for polynomials: except for trivial cases, An + Bn + Cn = 0 implies n ≤ 2 for coprime polynomials A, B, C over a field.

[FLP] title = A Constructive Proof for FLP author = Benjamin Bisping , Paul-David Brodmann , Tim Jungnickel , Christina Rickmann , Henning Seidler , Anke Stüber , Arno Wilhelm-Weidner , Kirstin Peters , Uwe Nestmann date = 2016-05-18 topic = Computer science/Concurrency abstract = The impossibility of distributed consensus with one faulty process is a result with important consequences for real world distributed systems e.g., commits in replicated databases. Since proofs are not immune to faults and even plausible proofs with a profound formalism can conclude wrong results, we validate the fundamental result named FLP after Fischer, Lynch and Paterson. We present a formalization of distributed systems and the aforementioned consensus problem. Our proof is based on Hagen Völzer's paper "A constructive proof for FLP". In addition to the enhanced confidence in the validity of Völzer's proof, we contribute the missing gaps to show the correctness in Isabelle/HOL. We clarify the proof details and even prove fairness of the infinite execution that contradicts consensus. Our Isabelle formalization can also be reused for further proofs of properties of distributed systems. notify = henning.seidler@mailbox.tu-berlin.de [IMAP-CRDT] title = The IMAP CmRDT author = Tim Jungnickel , Lennart Oldenburg <>, Matthias Loibl <> topic = Computer science/Algorithms/Distributed, Computer science/Data structures date = 2017-11-09 notify = tim.jungnickel@tu-berlin.de abstract = We provide our Isabelle/HOL formalization of a Conflict-free Replicated Datatype for Internet Message Access Protocol commands. We show that Strong Eventual Consistency (SEC) is guaranteed by proving the commutativity of concurrent operations. We base our formalization on the recently proposed "framework for establishing Strong Eventual Consistency for Conflict-free Replicated Datatypes" (AFP.CRDT) from Gomes et al. Hence, we provide an additional example of how the recently proposed framework can be used to design and prove CRDTs. [Incredible_Proof_Machine] title = The meta theory of the Incredible Proof Machine author = Joachim Breitner , Denis Lohner date = 2016-05-20 topic = Logic/Proof theory abstract = The Incredible Proof Machine is an interactive visual theorem prover which represents proofs as port graphs. We model this proof representation in Isabelle, and prove that it is just as powerful as natural deduction. notify = mail@joachim-breitner.de [Word_Lib] title = Finite Machine Word Library author = Joel Beeren<>, Matthew Fernandez<>, Xin Gao<>, Gerwin Klein , Rafal Kolanski<>, Japheth Lim<>, Corey Lewis<>, Daniel Matichuk<>, Thomas Sewell<> notify = kleing@unsw.edu.au date = 2016-06-09 topic = Computer science/Data structures abstract = This entry contains an extension to the Isabelle library for fixed-width machine words. In particular, the entry adds quickcheck setup for words, printing as hexadecimals, additional operations, reasoning about alignment, signed words, enumerations of words, normalisation of word numerals, and an extensive library of properties about generic fixed-width words, as well as an instantiation of many of these to the commonly used 32 and 64-bit bases. [Catalan_Numbers] title = Catalan Numbers author = Manuel Eberl notify = manuel@pruvisto.org date = 2016-06-21 topic = Mathematics/Combinatorics abstract =

In this work, we define the Catalan numbers Cn and prove several equivalent definitions (including some closed-form formulae). We also show one of their applications (counting the number of binary trees of size n), prove the asymptotic growth approximation Cn ∼ 4n / (√π · n1.5), and provide reasonably efficient executable code to compute them.

The derivation of the closed-form formulae uses algebraic manipulations of the ordinary generating function of the Catalan numbers, and the asymptotic approximation is then done using generalised binomial coefficients and the Gamma function. Thanks to these highly non-elementary mathematical tools, the proofs are very short and simple.

[Fisher_Yates] title = Fisher–Yates shuffle author = Manuel Eberl notify = manuel@pruvisto.org date = 2016-09-30 topic = Computer science/Algorithms abstract =

This work defines and proves the correctness of the Fisher–Yates algorithm for shuffling – i.e. producing a random permutation – of a list. The algorithm proceeds by traversing the list and in each step swapping the current element with a random element from the remaining list.

[Bertrands_Postulate] title = Bertrand's postulate author = Julian Biendarra<>, Manuel Eberl contributors = Lawrence C. Paulson topic = Mathematics/Number theory date = 2017-01-17 notify = manuel@pruvisto.org abstract =

Bertrand's postulate is an early result on the distribution of prime numbers: For every positive integer n, there exists a prime number that lies strictly between n and 2n. The proof is ported from John Harrison's formalisation in HOL Light. It proceeds by first showing that the property is true for all n greater than or equal to 600 and then showing that it also holds for all n below 600 by case distinction.

[Rewriting_Z] title = The Z Property author = Bertram Felgenhauer<>, Julian Nagele<>, Vincent van Oostrom<>, Christian Sternagel notify = bertram.felgenhauer@uibk.ac.at, julian.nagele@uibk.ac.at, c.sternagel@gmail.com date = 2016-06-30 topic = Logic/Rewriting abstract = We formalize the Z property introduced by Dehornoy and van Oostrom. First we show that for any abstract rewrite system, Z implies confluence. Then we give two examples of proofs using Z: confluence of lambda-calculus with respect to beta-reduction and confluence of combinatory logic. [Resolution_FOL] title = The Resolution Calculus for First-Order Logic author = Anders Schlichtkrull notify = andschl@dtu.dk date = 2016-06-30 topic = Logic/General logic/Mechanization of proofs abstract = This theory is a formalization of the resolution calculus for first-order logic. It is proven sound and complete. The soundness proof uses the substitution lemma, which shows a correspondence between substitutions and updates to an environment. The completeness proof uses semantic trees, i.e. trees whose paths are partial Herbrand interpretations. It employs Herbrand's theorem in a formulation which states that an unsatisfiable set of clauses has a finite closed semantic tree. It also uses the lifting lemma which lifts resolution derivation steps from the ground world up to the first-order world. The theory is presented in a paper in the Journal of Automated Reasoning [Sch18] which extends a paper presented at the International Conference on Interactive Theorem Proving [Sch16]. An earlier version was presented in an MSc thesis [Sch15]. The formalization mostly follows textbooks by Ben-Ari [BA12], Chang and Lee [CL73], and Leitsch [Lei97]. The theory is part of the IsaFoL project [IsaFoL].

[Sch18] Anders Schlichtkrull. "Formalization of the Resolution Calculus for First-Order Logic". Journal of Automated Reasoning, 2018.
[Sch16] Anders Schlichtkrull. "Formalization of the Resolution Calculus for First-Order Logic". In: ITP 2016. Vol. 9807. LNCS. Springer, 2016.
[Sch15] Anders Schlichtkrull. "Formalization of Resolution Calculus in Isabelle". https://people.compute.dtu.dk/andschl/Thesis.pdf. MSc thesis. Technical University of Denmark, 2015.
[BA12] Mordechai Ben-Ari. Mathematical Logic for Computer Science. 3rd. Springer, 2012.
[CL73] Chin-Liang Chang and Richard Char-Tung Lee. Symbolic Logic and Mechanical Theorem Proving. 1st. Academic Press, Inc., 1973.
[Lei97] Alexander Leitsch. The Resolution Calculus. Texts in theoretical computer science. Springer, 1997.
[IsaFoL] IsaFoL authors. IsaFoL: Isabelle Formalization of Logic. https://bitbucket.org/jasmin_blanchette/isafol. extra-history = Change history: [2018-01-24]: added several new versions of the soundness and completeness theorems as described in the paper [Sch18].
[2018-03-20]: added a concrete instance of the unification and completeness theorems using the First-Order Terms AFP-entry from IsaFoR as described in the papers [Sch16] and [Sch18]. [Surprise_Paradox] title = Surprise Paradox author = Joachim Breitner notify = mail@joachim-breitner.de date = 2016-07-17 topic = Logic/Proof theory abstract = In 1964, Fitch showed that the paradox of the surprise hanging can be resolved by showing that the judge’s verdict is inconsistent. His formalization builds on Gödel’s coding of provability. In this theory, we reproduce his proof in Isabelle, building on Paulson’s formalisation of Gödel’s incompleteness theorems. [Ptolemys_Theorem] title = Ptolemy's Theorem author = Lukas Bulwahn notify = lukas.bulwahn@gmail.com date = 2016-08-07 topic = Mathematics/Geometry abstract = This entry provides an analytic proof to Ptolemy's Theorem using polar form transformation and trigonometric identities. In this formalization, we use ideas from John Harrison's HOL Light formalization and the proof sketch on the Wikipedia entry of Ptolemy's Theorem. This theorem is the 95th theorem of the Top 100 Theorems list. [Falling_Factorial_Sum] title = The Falling Factorial of a Sum author = Lukas Bulwahn topic = Mathematics/Combinatorics date = 2017-12-22 notify = lukas.bulwahn@gmail.com abstract = This entry shows that the falling factorial of a sum can be computed with an expression using binomial coefficients and the falling factorial of its summands. The entry provides three different proofs: a combinatorial proof, an induction proof and an algebraic proof using the Vandermonde identity. The three formalizations try to follow their informal presentations from a Mathematics Stack Exchange page as close as possible. The induction and algebraic formalization end up to be very close to their informal presentation, whereas the combinatorial proof first requires the introduction of list interleavings, and significant more detail than its informal presentation. [InfPathElimination] title = Infeasible Paths Elimination by Symbolic Execution Techniques: Proof of Correctness and Preservation of Paths author = Romain Aissat<>, Frederic Voisin<>, Burkhart Wolff notify = wolff@lri.fr date = 2016-08-18 topic = Computer science/Programming languages/Static analysis abstract = TRACER is a tool for verifying safety properties of sequential C programs. TRACER attempts at building a finite symbolic execution graph which over-approximates the set of all concrete reachable states and the set of feasible paths. We present an abstract framework for TRACER and similar CEGAR-like systems. The framework provides 1) a graph- transformation based method for reducing the feasible paths in control-flow graphs, 2) a model for symbolic execution, subsumption, predicate abstraction and invariant generation. In this framework we formally prove two key properties: correct construction of the symbolic states and preservation of feasible paths. The framework focuses on core operations, leaving to concrete prototypes to “fit in” heuristics for combining them. The accompanying paper (published in ITP 2016) can be found at https://www.lri.fr/∼wolff/papers/conf/2016-itp-InfPathsNSE.pdf. [Stirling_Formula] title = Stirling's formula author = Manuel Eberl notify = manuel@pruvisto.org date = 2016-09-01 topic = Mathematics/Analysis abstract =

This work contains a proof of Stirling's formula both for the factorial $n! \sim \sqrt{2\pi n} (n/e)^n$ on natural numbers and the real Gamma function $\Gamma(x)\sim \sqrt{2\pi/x} (x/e)^x$. The proof is based on work by Graham Jameson.

This is then extended to the full asymptotic expansion $$\log\Gamma(z) = \big(z - \tfrac{1}{2}\big)\log z - z + \tfrac{1}{2}\log(2\pi) + \sum_{k=1}^{n-1} \frac{B_{k+1}}{k(k+1)} z^{-k}\\ {} - \frac{1}{n} \int_0^\infty B_n([t])(t + z)^{-n}\,\text{d}t$$ uniformly for all complex $z\neq 0$ in the cone $\text{arg}(z)\leq \alpha$ for any $\alpha\in(0,\pi)$, with which the above asymptotic relation for Γ is also extended to complex arguments.

[Lp] title = Lp spaces author = Sebastien Gouezel notify = sebastien.gouezel@univ-rennes1.fr date = 2016-10-05 topic = Mathematics/Analysis abstract = Lp is the space of functions whose p-th power is integrable. It is one of the most fundamental Banach spaces that is used in analysis and probability. We develop a framework for function spaces, and then implement the Lp spaces in this framework using the existing integration theory in Isabelle/HOL. Our development contains most fundamental properties of Lp spaces, notably the Hölder and Minkowski inequalities, completeness of Lp, duality, stability under almost sure convergence, multiplication of functions in Lp and Lq, stability under conditional expectation. [Berlekamp_Zassenhaus] title = The Factorization Algorithm of Berlekamp and Zassenhaus author = Jose Divasón , Sebastiaan Joosten , René Thiemann , Akihisa Yamada notify = rene.thiemann@uibk.ac.at date = 2016-10-14 topic = Mathematics/Algebra abstract =

We formalize the Berlekamp-Zassenhaus algorithm for factoring square-free integer polynomials in Isabelle/HOL. We further adapt an existing formalization of Yun’s square-free factorization algorithm to integer polynomials, and thus provide an efficient and certified factorization algorithm for arbitrary univariate polynomials.

The algorithm first performs a factorization in the prime field GF(p) and then performs computations in the integer ring modulo p^k, where both p and k are determined at runtime. Since a natural modeling of these structures via dependent types is not possible in Isabelle/HOL, we formalize the whole algorithm using Isabelle’s recent addition of local type definitions.

Through experiments we verify that our algorithm factors polynomials of degree 100 within seconds.

[Allen_Calculus] title = Allen's Interval Calculus author = Fadoua Ghourabi <> notify = fadouaghourabi@gmail.com date = 2016-09-29 topic = Logic/General logic/Temporal logic, Mathematics/Order abstract = Allen’s interval calculus is a qualitative temporal representation of time events. Allen introduced 13 binary relations that describe all the possible arrangements between two events, i.e. intervals with non-zero finite length. The compositions are pertinent to reasoning about knowledge of time. In particular, a consistency problem of relation constraints is commonly solved with a guideline from these compositions. We formalize the relations together with an axiomatic system. We proof the validity of the 169 compositions of these relations. We also define nests as the sets of intervals that share a meeting point. We prove that nests give the ordering properties of points without introducing a new datatype for points. [1] J.F. Allen. Maintaining Knowledge about Temporal Intervals. In Commun. ACM, volume 26, pages 832–843, 1983. [2] J. F. Allen and P. J. Hayes. A Common-sense Theory of Time. In Proceedings of the 9th International Joint Conference on Artificial Intelligence (IJCAI’85), pages 528–531, 1985. [Source_Coding_Theorem] title = Source Coding Theorem author = Quentin Hibon , Lawrence C. Paulson notify = qh225@cl.cam.ac.uk date = 2016-10-19 topic = Mathematics/Probability theory abstract = This document contains a proof of the necessary condition on the code rate of a source code, namely that this code rate is bounded by the entropy of the source. This represents one half of Shannon's source coding theorem, which is itself an equivalence. [Buffons_Needle] title = Buffon's Needle Problem author = Manuel Eberl topic = Mathematics/Probability theory, Mathematics/Geometry date = 2017-06-06 notify = manuel@pruvisto.org abstract = In the 18th century, Georges-Louis Leclerc, Comte de Buffon posed and later solved the following problem, which is often called the first problem ever solved in geometric probability: Given a floor divided into vertical strips of the same width, what is the probability that a needle thrown onto the floor randomly will cross two strips? This entry formally defines the problem in the case where the needle's position is chosen uniformly at random in a single strip around the origin (which is equivalent to larger arrangements due to symmetry). It then provides proofs of the simple solution in the case where the needle's length is no greater than the width of the strips and the more complicated solution in the opposite case. [SPARCv8] title = A formal model for the SPARCv8 ISA and a proof of non-interference for the LEON3 processor author = Zhe Hou , David Sanan , Alwen Tiu , Yang Liu notify = zhe.hou@ntu.edu.sg, sanan@ntu.edu.sg date = 2016-10-19 topic = Computer science/Security, Computer science/Hardware abstract = We formalise the SPARCv8 instruction set architecture (ISA) which is used in processors such as LEON3. Our formalisation can be specialised to any SPARCv8 CPU, here we use LEON3 as a running example. Our model covers the operational semantics for all the instructions in the integer unit of the SPARCv8 architecture and it supports Isabelle code export, which effectively turns the Isabelle model into a SPARCv8 CPU simulator. We prove the language-based non-interference property for the LEON3 processor. Our model is based on deterministic monad, which is a modified version of the non-deterministic monad from NICTA/l4v. [Separata] title = Separata: Isabelle tactics for Separation Algebra author = Zhe Hou , David Sanan , Alwen Tiu , Rajeev Gore , Ranald Clouston notify = zhe.hou@ntu.edu.sg date = 2016-11-16 topic = Computer science/Programming languages/Logics, Tools abstract = We bring the labelled sequent calculus $LS_{PASL}$ for propositional abstract separation logic to Isabelle. The tactics given here are directly applied on an extension of the Separation Algebra in the AFP. In addition to the cancellative separation algebra, we further consider some useful properties in the heap model of separation logic, such as indivisible unit, disjointness, and cross-split. The tactics are essentially a proof search procedure for the calculus $LS_{PASL}$. We wrap the tactics in an Isabelle method called separata, and give a few examples of separation logic formulae which are provable by separata. [LOFT] title = LOFT — Verified Migration of Linux Firewalls to SDN author = Julius Michaelis , Cornelius Diekmann notify = isabelleopenflow@liftm.de date = 2016-10-21 topic = Computer science/Networks abstract = We present LOFT — Linux firewall OpenFlow Translator, a system that transforms the main routing table and FORWARD chain of iptables of a Linux-based firewall into a set of static OpenFlow rules. Our implementation is verified against a model of a simplified Linux-based router and we can directly show how much of the original functionality is preserved. [Stable_Matching] title = Stable Matching author = Peter Gammie notify = peteg42@gmail.com date = 2016-10-24 topic = Mathematics/Games and economics abstract = We mechanize proofs of several results from the matching with contracts literature, which generalize those of the classical two-sided matching scenarios that go by the name of stable marriage. Our focus is on game theoretic issues. Along the way we develop executable algorithms for computing optimal stable matches. [Modal_Logics_for_NTS] title = Modal Logics for Nominal Transition Systems author = Tjark Weber , Lars-Henrik Eriksson , Joachim Parrow , Johannes Borgström , Ramunas Gutkovas notify = tjark.weber@it.uu.se date = 2016-10-25 topic = Computer science/Concurrency/Process calculi, Logic/General logic/Modal logic abstract = We formalize a uniform semantic substrate for a wide variety of process calculi where states and action labels can be from arbitrary nominal sets. A Hennessy-Milner logic for these systems is defined, and proved adequate for bisimulation equivalence. A main novelty is the construction of an infinitary nominal data type to model formulas with (finitely supported) infinite conjunctions and actions that may contain binding names. The logic is generalized to treat different bisimulation variants such as early, late and open in a systematic way. extra-history = Change history: [2017-01-29]: Formalization of weak bisimilarity added (revision c87cc2057d9c) [Abs_Int_ITP2012] title = Abstract Interpretation of Annotated Commands author = Tobias Nipkow notify = nipkow@in.tum.de date = 2016-11-23 topic = Computer science/Programming languages/Static analysis abstract = This is the Isabelle formalization of the material decribed in the eponymous ITP 2012 paper. It develops a generic abstract interpreter for a while-language, including widening and narrowing. The collecting semantics and the abstract interpreter operate on annotated commands: the program is represented as a syntax tree with the semantic information directly embedded, without auxiliary labels. The aim of the formalization is simplicity, not efficiency or precision. This is motivated by the inclusion of the material in a theorem prover based course on semantics. A similar (but more polished) development is covered in the book Concrete Semantics. [Complx] title = COMPLX: A Verification Framework for Concurrent Imperative Programs author = Sidney Amani<>, June Andronick<>, Maksym Bortin<>, Corey Lewis<>, Christine Rizkallah<>, Joseph Tuong<> notify = sidney.amani@data61.csiro.au, corey.lewis@data61.csiro.au date = 2016-11-29 topic = Computer science/Programming languages/Logics, Computer science/Programming languages/Language definitions abstract = We propose a concurrency reasoning framework for imperative programs, based on the Owicki-Gries (OG) foundational shared-variable concurrency method. Our framework combines the approaches of Hoare-Parallel, a formalisation of OG in Isabelle/HOL for a simple while-language, and Simpl, a generic imperative language embedded in Isabelle/HOL, allowing formal reasoning on C programs. We define the Complx language, extending the syntax and semantics of Simpl with support for parallel composition and synchronisation. We additionally define an OG logic, which we prove sound w.r.t. the semantics, and a verification condition generator, both supporting involved low-level imperative constructs such as function calls and abrupt termination. We illustrate our framework on an example that features exceptions, guards and function calls. We aim to then target concurrent operating systems, such as the interruptible eChronos embedded operating system for which we already have a model-level OG proof using Hoare-Parallel. extra-history = Change history: [2017-01-13]: Improve VCG for nested parallels and sequential sections (revision 30739dbc3dcb) [Paraconsistency] title = Paraconsistency author = Anders Schlichtkrull , Jørgen Villadsen topic = Logic/General logic/Paraconsistent logics date = 2016-12-07 notify = andschl@dtu.dk, jovi@dtu.dk abstract = Paraconsistency is about handling inconsistency in a coherent way. In classical and intuitionistic logic everything follows from an inconsistent theory. A paraconsistent logic avoids the explosion. Quite a few applications in computer science and engineering are discussed in the Intelligent Systems Reference Library Volume 110: Towards Paraconsistent Engineering (Springer 2016). We formalize a paraconsistent many-valued logic that we motivated and described in a special issue on logical approaches to paraconsistency (Journal of Applied Non-Classical Logics 2005). We limit ourselves to the propositional fragment of the higher-order logic. The logic is based on so-called key equalities and has a countably infinite number of truth values. We prove theorems in the logic using the definition of validity. We verify truth tables and also counterexamples for non-theorems. We prove meta-theorems about the logic and finally we investigate a case study. [Proof_Strategy_Language] title = Proof Strategy Language author = Yutaka Nagashima<> topic = Tools date = 2016-12-20 notify = Yutaka.Nagashima@data61.csiro.au abstract = Isabelle includes various automatic tools for finding proofs under certain conditions. However, for each conjecture, knowing which automation to use, and how to tweak its parameters, is currently labour intensive. We have developed a language, PSL, designed to capture high level proof strategies. PSL offloads the construction of human-readable fast-to-replay proof scripts to automatic search, making use of search-time information about each conjecture. Our preliminary evaluations show that PSL reduces the labour cost of interactive theorem proving. This submission contains the implementation of PSL and an example theory file, Example.thy, showing how to write poof strategies in PSL. [Concurrent_Ref_Alg] title = Concurrent Refinement Algebra and Rely Quotients author = Julian Fell , Ian J. Hayes , Andrius Velykis topic = Computer science/Concurrency date = 2016-12-30 notify = Ian.Hayes@itee.uq.edu.au abstract = The concurrent refinement algebra developed here is designed to provide a foundation for rely/guarantee reasoning about concurrent programs. The algebra builds on a complete lattice of commands by providing sequential composition, parallel composition and a novel weak conjunction operator. The weak conjunction operator coincides with the lattice supremum providing its arguments are non-aborting, but aborts if either of its arguments do. Weak conjunction provides an abstract version of a guarantee condition as a guarantee process. We distinguish between models that distribute sequential composition over non-deterministic choice from the left (referred to as being conjunctive in the refinement calculus literature) and those that don't. Least and greatest fixed points of monotone functions are provided to allow recursion and iteration operators to be added to the language. Additional iteration laws are available for conjunctive models. The rely quotient of processes c and i is the process that, if executed in parallel with i implements c. It represents an abstract version of a rely condition generalised to a process. [FOL_Harrison] title = First-Order Logic According to Harrison author = Alexander Birch Jensen , Anders Schlichtkrull , Jørgen Villadsen topic = Logic/General logic/Mechanization of proofs date = 2017-01-01 notify = aleje@dtu.dk, andschl@dtu.dk, jovi@dtu.dk abstract =

We present a certified declarative first-order prover with equality based on John Harrison's Handbook of Practical Logic and Automated Reasoning, Cambridge University Press, 2009. ML code reflection is used such that the entire prover can be executed within Isabelle as a very simple interactive proof assistant. As examples we consider Pelletier's problems 1-46.

Reference: Programming and Verifying a Declarative First-Order Prover in Isabelle/HOL. Alexander Birch Jensen, John Bruntse Larsen, Anders Schlichtkrull & Jørgen Villadsen. AI Communications 31:281-299 2018. https://content.iospress.com/articles/ai-communications/aic764

See also: Students' Proof Assistant (SPA). https://github.com/logic-tools/spa

extra-history = Change history: [2018-07-21]: Proof of Pelletier's problem 34 (Andrews's Challenge) thanks to Asta Halkjær From. [Bernoulli] title = Bernoulli Numbers author = Lukas Bulwahn, Manuel Eberl topic = Mathematics/Analysis, Mathematics/Number theory date = 2017-01-24 notify = manuel@pruvisto.org abstract =

Bernoulli numbers were first discovered in the closed-form expansion of the sum 1m + 2m + … + nm for a fixed m and appear in many other places. This entry provides three different definitions for them: a recursive one, an explicit one, and one through their exponential generating function.

In addition, we prove some basic facts, e.g. their relation to sums of powers of integers and that all odd Bernoulli numbers except the first are zero, and some advanced facts like their relationship to the Riemann zeta function on positive even integers.

We also prove the correctness of the Akiyama–Tanigawa algorithm for computing Bernoulli numbers with reasonable efficiency, and we define the periodic Bernoulli polynomials (which appear e.g. in the Euler–MacLaurin summation formula and the expansion of the log-Gamma function) and prove their basic properties.

[Stone_Relation_Algebras] title = Stone Relation Algebras author = Walter Guttmann topic = Mathematics/Algebra date = 2017-02-07 notify = walter.guttmann@canterbury.ac.nz abstract = We develop Stone relation algebras, which generalise relation algebras by replacing the underlying Boolean algebra structure with a Stone algebra. We show that finite matrices over extended real numbers form an instance. As a consequence, relation-algebraic concepts and methods can be used for reasoning about weighted graphs. We also develop a fixpoint calculus and apply it to compare different definitions of reflexive-transitive closures in semirings. extra-history = Change history: [2017-07-05]: generalised extended reals to linear orders (revision b8e703159177) [Stone_Kleene_Relation_Algebras] title = Stone-Kleene Relation Algebras author = Walter Guttmann topic = Mathematics/Algebra date = 2017-07-06 notify = walter.guttmann@canterbury.ac.nz abstract = We develop Stone-Kleene relation algebras, which expand Stone relation algebras with a Kleene star operation to describe reachability in weighted graphs. Many properties of the Kleene star arise as a special case of a more general theory of iteration based on Conway semirings extended by simulation axioms. This includes several theorems representing complex program transformations. We formally prove the correctness of Conway's automata-based construction of the Kleene star of a matrix. We prove numerous results useful for reasoning about weighted graphs. [Abstract_Soundness] title = Abstract Soundness author = Jasmin Christian Blanchette , Andrei Popescu , Dmitriy Traytel topic = Logic/Proof theory date = 2017-02-10 notify = jasmin.blanchette@gmail.com abstract = A formalized coinductive account of the abstract development of Brotherston, Gorogiannis, and Petersen [APLAS 2012], in a slightly more general form since we work with arbitrary infinite proofs, which may be acyclic. This work is described in detail in an article by the authors, published in 2017 in the Journal of Automated Reasoning. The abstract proof can be instantiated for various formalisms, including first-order logic with inductive predicates. [Differential_Dynamic_Logic] title = Differential Dynamic Logic author = Brandon Bohrer topic = Logic/General logic/Modal logic, Computer science/Programming languages/Logics date = 2017-02-13 notify = bbohrer@cs.cmu.edu abstract = We formalize differential dynamic logic, a logic for proving properties of hybrid systems. The proof calculus in this formalization is based on the uniform substitution principle. We show it is sound with respect to our denotational semantics, which provides increased confidence in the correctness of the KeYmaera X theorem prover based on this calculus. As an application, we include a proof term checker embedded in Isabelle/HOL with several example proofs. Published in: Brandon Bohrer, Vincent Rahli, Ivana Vukotic, Marcus Völp, André Platzer: Formally verified differential dynamic logic. CPP 2017. [Syntax_Independent_Logic] title = Syntax-Independent Logic Infrastructure author = Andrei Popescu , Dmitriy Traytel topic = Logic/Proof theory date = 2020-09-16 notify = a.popescu@sheffield.ac.uk, traytel@di.ku.dk abstract = We formalize a notion of logic whose terms and formulas are kept abstract. In particular, logical connectives, substitution, free variables, and provability are not defined, but characterized by their general properties as locale assumptions. Based on this abstract characterization, we develop further reusable reasoning infrastructure. For example, we define parallel substitution (along with proving its characterizing theorems) from single-point substitution. Similarly, we develop a natural deduction style proof system starting from the abstract Hilbert-style one. These one-time efforts benefit different concrete logics satisfying our locales' assumptions. We instantiate the syntax-independent logic infrastructure to Robinson arithmetic (also known as Q) in the AFP entry Robinson_Arithmetic and to hereditarily finite set theory in the AFP entries Goedel_HFSet_Semantic and Goedel_HFSet_Semanticless, which are part of our formalization of Gödel's Incompleteness Theorems described in our CADE-27 paper A Formally Verified Abstract Account of Gödel's Incompleteness Theorems. [Goedel_Incompleteness] title = An Abstract Formalization of Gödel's Incompleteness Theorems author = Andrei Popescu , Dmitriy Traytel topic = Logic/Proof theory date = 2020-09-16 notify = a.popescu@sheffield.ac.uk, traytel@di.ku.dk abstract = We present an abstract formalization of Gödel's incompleteness theorems. We analyze sufficient conditions for the theorems' applicability to a partially specified logic. Our abstract perspective enables a comparison between alternative approaches from the literature. These include Rosser's variation of the first theorem, Jeroslow's variation of the second theorem, and the Swierczkowski–Paulson semantics-based approach. This AFP entry is the main entry point to the results described in our CADE-27 paper A Formally Verified Abstract Account of Gödel's Incompleteness Theorems. As part of our abstract formalization's validation, we instantiate our locales twice in the separate AFP entries Goedel_HFSet_Semantic and Goedel_HFSet_Semanticless. [Goedel_HFSet_Semantic] title = From Abstract to Concrete Gödel's Incompleteness Theorems—Part I author = Andrei Popescu , Dmitriy Traytel topic = Logic/Proof theory date = 2020-09-16 notify = a.popescu@sheffield.ac.uk, traytel@di.ku.dk abstract = We validate an abstract formulation of Gödel's First and Second Incompleteness Theorems from a separate AFP entry by instantiating them to the case of finite sound extensions of the Hereditarily Finite (HF) Set theory, i.e., FOL theories extending the HF Set theory with a finite set of axioms that are sound in the standard model. The concrete results had been previously formalised in an AFP entry by Larry Paulson; our instantiation reuses the infrastructure developed in that entry. [Goedel_HFSet_Semanticless] title = From Abstract to Concrete Gödel's Incompleteness Theorems—Part II author = Andrei Popescu , Dmitriy Traytel topic = Logic/Proof theory date = 2020-09-16 notify = a.popescu@sheffield.ac.uk, traytel@di.ku.dk abstract = We validate an abstract formulation of Gödel's Second Incompleteness Theorem from a separate AFP entry by instantiating it to the case of finite consistent extensions of the Hereditarily Finite (HF) Set theory, i.e., consistent FOL theories extending the HF Set theory with a finite set of axioms. The instantiation draws heavily on infrastructure previously developed by Larry Paulson in his direct formalisation of the concrete result. It strengthens Paulson's formalization of Gödel's Second from that entry by not assuming soundness, and in fact not relying on any notion of model or semantic interpretation. The strengthening was obtained by first replacing some of Paulson’s semantic arguments with proofs within his HF calculus, and then plugging in some of Paulson's (modified) lemmas to instantiate our soundness-free Gödel's Second locale. [Robinson_Arithmetic] title = Robinson Arithmetic author = Andrei Popescu , Dmitriy Traytel topic = Logic/Proof theory date = 2020-09-16 notify = a.popescu@sheffield.ac.uk, traytel@di.ku.dk abstract = We instantiate our syntax-independent logic infrastructure developed in a separate AFP entry to the FOL theory of Robinson arithmetic (also known as Q). The latter was formalised using Nominal Isabelle by adapting Larry Paulson’s formalization of the Hereditarily Finite Set theory. [Elliptic_Curves_Group_Law] title = The Group Law for Elliptic Curves author = Stefan Berghofer topic = Computer science/Security/Cryptography date = 2017-02-28 notify = berghofe@in.tum.de abstract = We prove the group law for elliptic curves in Weierstrass form over fields of characteristic greater than 2. In addition to affine coordinates, we also formalize projective coordinates, which allow for more efficient computations. By specializing the abstract formalization to prime fields, we can apply the curve operations to parameters used in standard security protocols. [Example-Submission] title = Example Submission author = Gerwin Klein topic = Mathematics/Analysis, Mathematics/Number theory date = 2004-02-25 notify = kleing@cse.unsw.edu.au abstract =

This is an example submission to the Archive of Formal Proofs. It shows submission requirements and explains the structure of a simple typical submission.

Note that you can use HTML tags and LaTeX formulae like $\sum_{n=1}^\infty \frac{1}{n^2} = \frac{\pi^2}{6}$ in the abstract. Display formulae like $$ \int_0^1 x^{-x}\,\text{d}x = \sum_{n=1}^\infty n^{-n}$$ are also possible. Please read the submission guidelines before using this.

extra-no-index = no-index: true [CRDT] title = A framework for establishing Strong Eventual Consistency for Conflict-free Replicated Datatypes author = Victor B. F. Gomes , Martin Kleppmann, Dominic P. Mulligan, Alastair R. Beresford topic = Computer science/Algorithms/Distributed, Computer science/Data structures date = 2017-07-07 notify = vb358@cam.ac.uk, dominic.p.mulligan@googlemail.com abstract = In this work, we focus on the correctness of Conflict-free Replicated Data Types (CRDTs), a class of algorithm that provides strong eventual consistency guarantees for replicated data. We develop a modular and reusable framework for verifying the correctness of CRDT algorithms. We avoid correctness issues that have dogged previous mechanised proofs in this area by including a network model in our formalisation, and proving that our theorems hold in all possible network behaviours. Our axiomatic network model is a standard abstraction that accurately reflects the behaviour of real-world computer networks. Moreover, we identify an abstract convergence theorem, a property of order relations, which provides a formal definition of strong eventual consistency. We then obtain the first machine-checked correctness theorems for three concrete CRDTs: the Replicated Growable Array, the Observed-Remove Set, and an Increment-Decrement Counter. [HOLCF-Prelude] title = HOLCF-Prelude author = Joachim Breitner, Brian Huffman<>, Neil Mitchell<>, Christian Sternagel topic = Computer science/Functional programming date = 2017-07-15 notify = c.sternagel@gmail.com, joachim@cis.upenn.edu, hupel@in.tum.de abstract = The Isabelle/HOLCF-Prelude is a formalization of a large part of Haskell's standard prelude in Isabelle/HOLCF. We use it to prove the correctness of the Eratosthenes' Sieve, in its self-referential implementation commonly used to showcase Haskell's laziness; prove correctness of GHC's "fold/build" rule and related rewrite rules; and certify a number of hints suggested by HLint. [Decl_Sem_Fun_PL] title = Declarative Semantics for Functional Languages author = Jeremy Siek topic = Computer science/Programming languages date = 2017-07-21 notify = jsiek@indiana.edu abstract = We present a semantics for an applied call-by-value lambda-calculus that is compositional, extensional, and elementary. We present four different views of the semantics: 1) as a relational (big-step) semantics that is not operational but instead declarative, 2) as a denotational semantics that does not use domain theory, 3) as a non-deterministic interpreter, and 4) as a variant of the intersection type systems of the Torino group. We prove that the semantics is correct by showing that it is sound and complete with respect to operational semantics on programs and that is sound with respect to contextual equivalence. We have not yet investigated whether it is fully abstract. We demonstrate that this approach to semantics is useful with three case studies. First, we use the semantics to prove correctness of a compiler optimization that inlines function application. Second, we adapt the semantics to the polymorphic lambda-calculus extended with general recursion and prove semantic type soundness. Third, we adapt the semantics to the call-by-value lambda-calculus with mutable references.
The paper that accompanies these Isabelle theories is available on arXiv. [DynamicArchitectures] title = Dynamic Architectures author = Diego Marmsoler topic = Computer science/System description languages date = 2017-07-28 notify = diego.marmsoler@tum.de abstract = The architecture of a system describes the system's overall organization into components and connections between those components. With the emergence of mobile computing, dynamic architectures have become increasingly important. In such architectures, components may appear or disappear, and connections may change over time. In the following we mechanize a theory of dynamic architectures and verify the soundness of a corresponding calculus. Therefore, we first formalize the notion of configuration traces as a model for dynamic architectures. Then, the behavior of single components is formalized in terms of behavior traces and an operator is introduced and studied to extract the behavior of a single component out of a given configuration trace. Then, behavior trace assertions are introduced as a temporal specification technique to specify behavior of components. Reasoning about component behavior in a dynamic context is formalized in terms of a calculus for dynamic architectures. Finally, the soundness of the calculus is verified by introducing an alternative interpretation for behavior trace assertions over configuration traces and proving the rules of the calculus. Since projection may lead to finite as well as infinite behavior traces, they are formalized in terms of coinductive lists. Thus, our theory is based on Lochbihler's formalization of coinductive lists. The theory may be applied to verify properties for dynamic architectures. extra-history = Change history: [2018-06-07]: adding logical operators to specify configuration traces (revision 09178f08f050)
[Stewart_Apollonius] title = Stewart's Theorem and Apollonius' Theorem author = Lukas Bulwahn topic = Mathematics/Geometry date = 2017-07-31 notify = lukas.bulwahn@gmail.com abstract = This entry formalizes the two geometric theorems, Stewart's and Apollonius' theorem. Stewart's Theorem relates the length of a triangle's cevian to the lengths of the triangle's two sides. Apollonius' Theorem is a specialisation of Stewart's theorem, restricting the cevian to be the median. The proof applies the law of cosines, some basic geometric facts about triangles and then simply transforms the terms algebraically to yield the conjectured relation. The formalization in Isabelle can closely follow the informal proofs described in the Wikipedia articles of those two theorems. [LambdaMu] title = The LambdaMu-calculus author = Cristina Matache , Victor B. F. Gomes , Dominic P. Mulligan topic = Computer science/Programming languages/Lambda calculi, Logic/General logic/Lambda calculus date = 2017-08-16 notify = victorborgesfg@gmail.com, dominic.p.mulligan@googlemail.com abstract = The propositions-as-types correspondence is ordinarily presented as linking the metatheory of typed λ-calculi and the proof theory of intuitionistic logic. Griffin observed that this correspondence could be extended to classical logic through the use of control operators. This observation set off a flurry of further research, leading to the development of Parigots λμ-calculus. In this work, we formalise λμ- calculus in Isabelle/HOL and prove several metatheoretical properties such as type preservation and progress. [Orbit_Stabiliser] title = Orbit-Stabiliser Theorem with Application to Rotational Symmetries author = Jonas Rädle topic = Mathematics/Algebra date = 2017-08-20 notify = jonas.raedle@tum.de abstract = The Orbit-Stabiliser theorem is a basic result in the algebra of groups that factors the order of a group into the sizes of its orbits and stabilisers. We formalize the notion of a group action and the related concepts of orbits and stabilisers. This allows us to prove the orbit-stabiliser theorem. In the second part of this work, we formalize the tetrahedral group and use the orbit-stabiliser theorem to prove that there are twelve (orientation-preserving) rotations of the tetrahedron. [PLM] title = Representation and Partial Automation of the Principia Logico-Metaphysica in Isabelle/HOL author = Daniel Kirchner topic = Logic/Philosophical aspects date = 2017-09-17 notify = daniel@ekpyron.org abstract =

We present an embedding of the second-order fragment of the Theory of Abstract Objects as described in Edward Zalta's upcoming work Principia Logico-Metaphysica (PLM) in the automated reasoning framework Isabelle/HOL. The Theory of Abstract Objects is a metaphysical theory that reifies property patterns, as they for example occur in the abstract reasoning of mathematics, as abstract objects and provides an axiomatic framework that allows to reason about these objects. It thereby serves as a fundamental metaphysical theory that can be used to axiomatize and describe a wide range of philosophical objects, such as Platonic forms or Leibniz' concepts, and has the ambition to function as a foundational theory of mathematics. The target theory of our embedding as described in chapters 7-9 of PLM employs a modal relational type theory as logical foundation for which a representation in functional type theory is known to be challenging.

Nevertheless we arrive at a functioning representation of the theory in the functional logic of Isabelle/HOL based on a semantical representation of an Aczel-model of the theory. Based on this representation we construct an implementation of the deductive system of PLM which allows to automatically and interactively find and verify theorems of PLM.

Our work thereby supports the concept of shallow semantical embeddings of logical systems in HOL as a universal tool for logical reasoning as promoted by Christoph Benzmüller.

The most notable result of the presented work is the discovery of a previously unknown paradox in the formulation of the Theory of Abstract Objects. The embedding of the theory in Isabelle/HOL played a vital part in this discovery. Furthermore it was possible to immediately offer several options to modify the theory to guarantee its consistency. Thereby our work could provide a significant contribution to the development of a proper grounding for object theory.

[KD_Tree] title = Multidimensional Binary Search Trees author = Martin Rau<> topic = Computer science/Data structures date = 2019-05-30 notify = martin.rau@tum.de, mrtnrau@googlemail.com abstract = This entry provides a formalization of multidimensional binary trees, also known as k-d trees. It includes a balanced build algorithm as well as the nearest neighbor algorithm and the range search algorithm. It is based on the papers Multidimensional binary search trees used for associative searching and An Algorithm for Finding Best Matches in Logarithmic Expected Time. extra-history = Change history: [2020-15-04]: Change representation of k-dimensional points from 'list' to HOL-Analysis.Finite_Cartesian_Product 'vec'. Update proofs to incorporate HOL-Analysis 'dist' and 'cbox' primitives. [Closest_Pair_Points] title = Closest Pair of Points Algorithms author = Martin Rau , Tobias Nipkow topic = Computer science/Algorithms/Geometry date = 2020-01-13 notify = martin.rau@tum.de, nipkow@in.tum.de abstract = This entry provides two related verified divide-and-conquer algorithms solving the fundamental Closest Pair of Points problem in Computational Geometry. Functional correctness and the optimal running time of O(n log n) are proved. Executable code is generated which is empirically competitive with handwritten reference implementations. extra-history = Change history: [2020-14-04]: Incorporate Time_Monad of the AFP entry Root_Balanced_Tree. [Approximation_Algorithms] title = Verified Approximation Algorithms author = Robin Eßmann , Tobias Nipkow , Simon Robillard , Ujkan Sulejmani<> topic = Computer science/Algorithms/Approximation date = 2020-01-16 notify = nipkow@in.tum.de abstract = We present the first formal verification of approximation algorithms for NP-complete optimization problems: vertex cover, set cover, independent set, center selection, load balancing, and bin packing. The proofs correct incompletenesses in existing proofs and improve the approximation ratio in one case. A detailed description of our work (excluding center selection) has been published in the proceedings of IJCAR 2020. [Diophantine_Eqns_Lin_Hom] title = Homogeneous Linear Diophantine Equations author = Florian Messner , Julian Parsert , Jonas Schöpf , Christian Sternagel topic = Computer science/Algorithms/Mathematical, Mathematics/Number theory, Tools license = LGPL date = 2017-10-14 notify = c.sternagel@gmail.com, julian.parsert@gmail.com abstract = We formalize the theory of homogeneous linear diophantine equations, focusing on two main results: (1) an abstract characterization of minimal complete sets of solutions, and (2) an algorithm computing them. Both, the characterization and the algorithm are based on previous work by Huet. Our starting point is a simple but inefficient variant of Huet's lexicographic algorithm incorporating improved bounds due to Clausen and Fortenbacher. We proceed by proving its soundness and completeness. Finally, we employ code equations to obtain a reasonably efficient implementation. Thus, we provide a formally verified solver for homogeneous linear diophantine equations. [Winding_Number_Eval] title = Evaluate Winding Numbers through Cauchy Indices author = Wenda Li topic = Mathematics/Analysis date = 2017-10-17 notify = wl302@cam.ac.uk, liwenda1990@hotmail.com abstract = In complex analysis, the winding number measures the number of times a path (counterclockwise) winds around a point, while the Cauchy index can approximate how the path winds. This entry provides a formalisation of the Cauchy index, which is then shown to be related to the winding number. In addition, this entry also offers a tactic that enables users to evaluate the winding number by calculating Cauchy indices. [Count_Complex_Roots] title = Count the Number of Complex Roots author = Wenda Li topic = Mathematics/Analysis date = 2017-10-17 notify = wl302@cam.ac.uk, liwenda1990@hotmail.com abstract = Based on evaluating Cauchy indices through remainder sequences, this entry provides an effective procedure to count the number of complex roots (with multiplicity) of a polynomial within various shapes (e.g., rectangle, circle and half-plane). Potential applications of this entry include certified complex root isolation (of a polynomial) and testing the Routh-Hurwitz stability criterion (i.e., to check whether all the roots of some characteristic polynomial have negative real parts). extra-history = Change history: [2021-10-26]: resolved the roots-on-the-border problem in the rectangular case (revision 82a159e398cf). [Buchi_Complementation] title = Büchi Complementation author = Julian Brunner topic = Computer science/Automata and formal languages date = 2017-10-19 notify = brunnerj@in.tum.de abstract = This entry provides a verified implementation of rank-based Büchi Complementation. The verification is done in three steps:
  1. Definition of odd rankings and proof that an automaton rejects a word iff there exists an odd ranking for it.
  2. Definition of the complement automaton and proof that it accepts exactly those words for which there is an odd ranking.
  3. Verified implementation of the complement automaton using the Isabelle Collections Framework.
[Transition_Systems_and_Automata] title = Transition Systems and Automata author = Julian Brunner topic = Computer science/Automata and formal languages date = 2017-10-19 notify = brunnerj@in.tum.de abstract = This entry provides a very abstract theory of transition systems that can be instantiated to express various types of automata. A transition system is typically instantiated by providing a set of initial states, a predicate for enabled transitions, and a transition execution function. From this, it defines the concepts of finite and infinite paths as well as the set of reachable states, among other things. Many useful theorems, from basic path manipulation rules to coinduction and run construction rules, are proven in this abstract transition system context. The library comes with instantiations for DFAs, NFAs, and Büchi automata. [Kuratowski_Closure_Complement] title = The Kuratowski Closure-Complement Theorem author = Peter Gammie , Gianpaolo Gioiosa<> topic = Mathematics/Topology date = 2017-10-26 notify = peteg42@gmail.com abstract = We discuss a topological curiosity discovered by Kuratowski (1922): the fact that the number of distinct operators on a topological space generated by compositions of closure and complement never exceeds 14, and is exactly 14 in the case of R. In addition, we prove a theorem due to Chagrov (1982) that classifies topological spaces according to the number of such operators they support. [Hybrid_Multi_Lane_Spatial_Logic] title = Hybrid Multi-Lane Spatial Logic author = Sven Linker topic = Logic/General logic/Modal logic date = 2017-11-06 notify = s.linker@liverpool.ac.uk abstract = We present a semantic embedding of a spatio-temporal multi-modal logic, specifically defined to reason about motorway traffic, into Isabelle/HOL. The semantic model is an abstraction of a motorway, emphasising local spatial properties, and parameterised by the types of sensors deployed in the vehicles. We use the logic to define controller constraints to ensure safety, i.e., the absence of collisions on the motorway. After proving safety with a restrictive definition of sensors, we relax these assumptions and show how to amend the controller constraints to still guarantee safety. [Dirichlet_L] title = Dirichlet L-Functions and Dirichlet's Theorem author = Manuel Eberl topic = Mathematics/Number theory, Mathematics/Algebra date = 2017-12-21 notify = manuel@pruvisto.org abstract =

This article provides a formalisation of Dirichlet characters and Dirichlet L-functions including proofs of their basic properties – most notably their analyticity, their areas of convergence, and their non-vanishing for ℜ(s) ≥ 1. All of this is built in a very high-level style using Dirichlet series. The proof of the non-vanishing follows a very short and elegant proof by Newman, which we attempt to reproduce faithfully in a similar level of abstraction in Isabelle.

This also leads to a relatively short proof of Dirichlet’s Theorem, which states that, if h and n are coprime, there are infinitely many primes p with ph (mod n).

[Symmetric_Polynomials] title = Symmetric Polynomials author = Manuel Eberl topic = Mathematics/Algebra date = 2018-09-25 notify = manuel@pruvisto.org abstract =

A symmetric polynomial is a polynomial in variables X1,…,Xn that does not discriminate between its variables, i. e. it is invariant under any permutation of them. These polynomials are important in the study of the relationship between the coefficients of a univariate polynomial and its roots in its algebraic closure.

This article provides a definition of symmetric polynomials and the elementary symmetric polynomials e1,…,en and proofs of their basic properties, including three notable ones:

  • Vieta's formula, which gives an explicit expression for the k-th coefficient of a univariate monic polynomial in terms of its roots x1,…,xn, namely ck = (-1)n-k en-k(x1,…,xn).
  • Second, the Fundamental Theorem of Symmetric Polynomials, which states that any symmetric polynomial is itself a uniquely determined polynomial combination of the elementary symmetric polynomials.
  • Third, as a corollary of the previous two, that given a polynomial over some ring R, any symmetric polynomial combination of its roots is also in R even when the roots are not.

Both the symmetry property itself and the witness for the Fundamental Theorem are executable.

[Taylor_Models] title = Taylor Models author = Christoph Traut<>, Fabian Immler topic = Computer science/Algorithms/Mathematical, Computer science/Data structures, Mathematics/Analysis, Mathematics/Algebra date = 2018-01-08 notify = immler@in.tum.de abstract = We present a formally verified implementation of multivariate Taylor models. Taylor models are a form of rigorous polynomial approximation, consisting of an approximation polynomial based on Taylor expansions, combined with a rigorous bound on the approximation error. Taylor models were introduced as a tool to mitigate the dependency problem of interval arithmetic. Our implementation automatically computes Taylor models for the class of elementary functions, expressed by composition of arithmetic operations and basic functions like exp, sin, or square root. [Green] title = An Isabelle/HOL formalisation of Green's Theorem author = Mohammad Abdulaziz , Lawrence C. Paulson topic = Mathematics/Analysis date = 2018-01-11 notify = mohammad.abdulaziz8@gmail.com, lp15@cam.ac.uk abstract = We formalise a statement of Green’s theorem—the first formalisation to our knowledge—in Isabelle/HOL. The theorem statement that we formalise is enough for most applications, especially in physics and engineering. Our formalisation is made possible by a novel proof that avoids the ubiquitous line integral cancellation argument. This eliminates the need to formalise orientations and region boundaries explicitly with respect to the outwards-pointing normal vector. Instead we appeal to a homological argument about equivalences between paths. [AI_Planning_Languages_Semantics] title = AI Planning Languages Semantics author = Mohammad Abdulaziz , Peter Lammich topic = Computer science/Artificial intelligence date = 2020-10-29 notify = mohammad.abdulaziz8@gmail.com abstract = This is an Isabelle/HOL formalisation of the semantics of the multi-valued planning tasks language that is used by the planning system Fast-Downward, the STRIPS fragment of the Planning Domain Definition Language (PDDL), and the STRIPS soundness meta-theory developed by Vladimir Lifschitz. It also contains formally verified checkers for checking the well-formedness of problems specified in either language as well the correctness of potential solutions. The formalisation in this entry was described in an earlier publication. [Verified_SAT_Based_AI_Planning] title = Verified SAT-Based AI Planning author = Mohammad Abdulaziz , Friedrich Kurz <> topic = Computer science/Artificial intelligence date = 2020-10-29 notify = mohammad.abdulaziz8@gmail.com abstract = We present an executable formally verified SAT encoding of classical AI planning that is based on the encodings by Kautz and Selman and the one by Rintanen et al. The encoding was experimentally tested and shown to be usable for reasonably sized standard AI planning benchmarks. We also use it as a reference to test a state-of-the-art SAT-based planner, showing that it sometimes falsely claims that problems have no solutions of certain lengths. The formalisation in this submission was described in an independent publication. [Gromov_Hyperbolicity] title = Gromov Hyperbolicity author = Sebastien Gouezel<> topic = Mathematics/Geometry date = 2018-01-16 notify = sebastien.gouezel@univ-rennes1.fr abstract = A geodesic metric space is Gromov hyperbolic if all its geodesic triangles are thin, i.e., every side is contained in a fixed thickening of the two other sides. While this definition looks innocuous, it has proved extremely important and versatile in modern geometry since its introduction by Gromov. We formalize the basic classical properties of Gromov hyperbolic spaces, notably the Morse lemma asserting that quasigeodesics are close to geodesics, the invariance of hyperbolicity under quasi-isometries, we define and study the Gromov boundary and its associated distance, and prove that a quasi-isometry between Gromov hyperbolic spaces extends to a homeomorphism of the boundaries. We also prove a less classical theorem, by Bonk and Schramm, asserting that a Gromov hyperbolic space embeds isometrically in a geodesic Gromov-hyperbolic space. As the original proof uses a transfinite sequence of Cauchy completions, this is an interesting formalization exercise. Along the way, we introduce basic material on isometries, quasi-isometries, Lipschitz maps, geodesic spaces, the Hausdorff distance, the Cauchy completion of a metric space, and the exponential on extended real numbers. [Ordered_Resolution_Prover] title = Formalization of Bachmair and Ganzinger's Ordered Resolution Prover author = Anders Schlichtkrull , Jasmin Christian Blanchette , Dmitriy Traytel , Uwe Waldmann topic = Logic/General logic/Mechanization of proofs date = 2018-01-18 notify = andschl@dtu.dk, j.c.blanchette@vu.nl abstract = This Isabelle/HOL formalization covers Sections 2 to 4 of Bachmair and Ganzinger's "Resolution Theorem Proving" chapter in the Handbook of Automated Reasoning. This includes soundness and completeness of unordered and ordered variants of ground resolution with and without literal selection, the standard redundancy criterion, a general framework for refutational theorem proving, and soundness and completeness of an abstract first-order prover. [Chandy_Lamport] title = A Formal Proof of The Chandy--Lamport Distributed Snapshot Algorithm author = Ben Fiedler , Dmitriy Traytel topic = Computer science/Algorithms/Distributed date = 2020-07-21 notify = ben.fiedler@inf.ethz.ch, traytel@inf.ethz.ch abstract = We provide a suitable distributed system model and implementation of the Chandy--Lamport distributed snapshot algorithm [ACM Transactions on Computer Systems, 3, 63-75, 1985]. Our main result is a formal termination and correctness proof of the Chandy--Lamport algorithm and its use in stable property detection. [BNF_Operations] title = Operations on Bounded Natural Functors author = Jasmin Christian Blanchette , Andrei Popescu , Dmitriy Traytel topic = Tools date = 2017-12-19 notify = jasmin.blanchette@gmail.com,uuomul@yahoo.com,traytel@inf.ethz.ch abstract = This entry formalizes the closure property of bounded natural functors (BNFs) under seven operations. These operations and the corresponding proofs constitute the core of Isabelle's (co)datatype package. To be close to the implemented tactics, the proofs are deliberately formulated as detailed apply scripts. The (co)datatypes together with (co)induction principles and (co)recursors are byproducts of the fixpoint operations LFP and GFP. Composition of BNFs is subdivided into four simpler operations: Compose, Kill, Lift, and Permute. The N2M operation provides mutual (co)induction principles and (co)recursors for nested (co)datatypes. [LLL_Basis_Reduction] title = A verified LLL algorithm author = Ralph Bottesch <>, Jose Divasón , Maximilian Haslbeck , Sebastiaan Joosten , René Thiemann , Akihisa Yamada<> topic = Computer science/Algorithms/Mathematical, Mathematics/Algebra date = 2018-02-02 notify = ralph.bottesch@uibk.ac.at, jose.divason@unirioja.es, maximilian.haslbeck@uibk.ac.at, s.j.c.joosten@utwente.nl, rene.thiemann@uibk.ac.at, ayamada@trs.cm.is.nagoya-u.ac.jp abstract = The Lenstra-Lenstra-Lovász basis reduction algorithm, also known as LLL algorithm, is an algorithm to find a basis with short, nearly orthogonal vectors of an integer lattice. Thereby, it can also be seen as an approximation to solve the shortest vector problem (SVP), which is an NP-hard problem, where the approximation quality solely depends on the dimension of the lattice, but not the lattice itself. The algorithm also possesses many applications in diverse fields of computer science, from cryptanalysis to number theory, but it is specially well-known since it was used to implement the first polynomial-time algorithm to factor polynomials. In this work we present the first mechanized soundness proof of the LLL algorithm to compute short vectors in lattices. The formalization follows a textbook by von zur Gathen and Gerhard. extra-history = Change history: [2018-04-16]: Integrated formal complexity bounds (Haslbeck, Thiemann) [2018-05-25]: Integrated much faster LLL implementation based on integer arithmetic (Bottesch, Haslbeck, Thiemann) [LLL_Factorization] title = A verified factorization algorithm for integer polynomials with polynomial complexity author = Jose Divasón , Sebastiaan Joosten , René Thiemann , Akihisa Yamada topic = Mathematics/Algebra date = 2018-02-06 notify = jose.divason@unirioja.es, s.j.c.joosten@utwente.nl, rene.thiemann@uibk.ac.at, ayamada@trs.cm.is.nagoya-u.ac.jp abstract = Short vectors in lattices and factors of integer polynomials are related. Each factor of an integer polynomial belongs to a certain lattice. When factoring polynomials, the condition that we are looking for an irreducible polynomial means that we must look for a small element in a lattice, which can be done by a basis reduction algorithm. In this development we formalize this connection and thereby one main application of the LLL basis reduction algorithm: an algorithm to factor square-free integer polynomials which runs in polynomial time. The work is based on our previous Berlekamp–Zassenhaus development, where the exponential reconstruction phase has been replaced by the polynomial-time basis reduction algorithm. Thanks to this formalization we found a serious flaw in a textbook. [Treaps] title = Treaps author = Maximilian Haslbeck , Manuel Eberl , Tobias Nipkow topic = Computer science/Data structures date = 2018-02-06 notify = manuel@pruvisto.org abstract =

A Treap is a binary tree whose nodes contain pairs consisting of some payload and an associated priority. It must have the search-tree property w.r.t. the payloads and the heap property w.r.t. the priorities. Treaps are an interesting data structure that is related to binary search trees (BSTs) in the following way: if one forgets all the priorities of a treap, the resulting BST is exactly the same as if one had inserted the elements into an empty BST in order of ascending priority. This means that a treap behaves like a BST where we can pretend the elements were inserted in a different order from the one in which they were actually inserted.

In particular, by choosing these priorities at random upon insertion of an element, we can pretend that we inserted the elements in random order, so that the shape of the resulting tree is that of a random BST no matter in what order we insert the elements. This is the main result of this formalisation.

[Skip_Lists] title = Skip Lists author = Max W. Haslbeck , Manuel Eberl topic = Computer science/Data structures date = 2020-01-09 notify = max.haslbeck@gmx.de abstract =

Skip lists are sorted linked lists enhanced with shortcuts and are an alternative to binary search trees. A skip lists consists of multiple levels of sorted linked lists where a list on level n is a subsequence of the list on level n − 1. In the ideal case, elements are skipped in such a way that a lookup in a skip lists takes O(log n) time. In a randomised skip list the skipped elements are choosen randomly.

This entry contains formalized proofs of the textbook results about the expected height and the expected length of a search path in a randomised skip list.

[Mersenne_Primes] title = Mersenne primes and the Lucas–Lehmer test author = Manuel Eberl topic = Mathematics/Number theory date = 2020-01-17 notify = manuel@pruvisto.org abstract =

This article provides formal proofs of basic properties of Mersenne numbers, i. e. numbers of the form 2n - 1, and especially of Mersenne primes.

In particular, an efficient, verified, and executable version of the Lucas–Lehmer test is developed. This test decides primality for Mersenne numbers in time polynomial in n.

[Hoare_Time] title = Hoare Logics for Time Bounds author = Maximilian P. L. Haslbeck , Tobias Nipkow topic = Computer science/Programming languages/Logics date = 2018-02-26 notify = haslbema@in.tum.de abstract = We study three different Hoare logics for reasoning about time bounds of imperative programs and formalize them in Isabelle/HOL: a classical Hoare like logic due to Nielson, a logic with potentials due to Carbonneaux et al. and a separation logic following work by Atkey, Chaguérand and Pottier. These logics are formally shown to be sound and complete. Verification condition generators are developed and are shown sound and complete too. We also consider variants of the systems where we abstract from multiplicative constants in the running time bounds, thus supporting a big-O style of reasoning. Finally we compare the expressive power of the three systems. [Architectural_Design_Patterns] title = A Theory of Architectural Design Patterns author = Diego Marmsoler topic = Computer science/System description languages date = 2018-03-01 notify = diego.marmsoler@tum.de abstract = The following document formalizes and verifies several architectural design patterns. Each pattern specification is formalized in terms of a locale where the locale assumptions correspond to the assumptions which a pattern poses on an architecture. Thus, pattern specifications may build on top of each other by interpreting the corresponding locale. A pattern is verified using the framework provided by the AFP entry Dynamic Architectures. Currently, the document consists of formalizations of 4 different patterns: the singleton, the publisher subscriber, the blackboard pattern, and the blockchain pattern. Thereby, the publisher component of the publisher subscriber pattern is modeled as an instance of the singleton pattern and the blackboard pattern is modeled as an instance of the publisher subscriber pattern. In general, this entry provides the first steps towards an overall theory of architectural design patterns. extra-history = Change history: [2018-05-25]: changing the major assumption for blockchain architectures from alternative minings to relative mining frequencies (revision 5043c5c71685)
[2019-04-08]: adapting the terminology: honest instead of trusted, dishonest instead of untrusted (revision 7af3431a22ae) [Weight_Balanced_Trees] title = Weight-Balanced Trees author = Tobias Nipkow , Stefan Dirix<> topic = Computer science/Data structures date = 2018-03-13 notify = nipkow@in.tum.de abstract = This theory provides a verified implementation of weight-balanced trees following the work of Hirai and Yamamoto who proved that all parameters in a certain range are valid, i.e. guarantee that insertion and deletion preserve weight-balance. Instead of a general theorem we provide parameterized proofs of preservation of the invariant that work for many (all?) valid parameters. [Fishburn_Impossibility] title = The Incompatibility of Fishburn-Strategyproofness and Pareto-Efficiency author = Felix Brandt , Manuel Eberl , Christian Saile , Christian Stricker topic = Mathematics/Games and economics date = 2018-03-22 notify = manuel@pruvisto.org abstract =

This formalisation contains the proof that there is no anonymous Social Choice Function for at least three agents and alternatives that fulfils both Pareto-Efficiency and Fishburn-Strategyproofness. It was derived from a proof of Brandt et al., which relies on an unverified translation of a fixed finite instance of the original problem to SAT. This Isabelle proof contains a machine-checked version of both the statement for exactly three agents and alternatives and the lifting to the general case.

[BNF_CC] title = Bounded Natural Functors with Covariance and Contravariance author = Andreas Lochbihler , Joshua Schneider topic = Computer science/Functional programming, Tools date = 2018-04-24 notify = mail@andreas-lochbihler.de, joshua.schneider@inf.ethz.ch abstract = Bounded natural functors (BNFs) provide a modular framework for the construction of (co)datatypes in higher-order logic. Their functorial operations, the mapper and relator, are restricted to a subset of the parameters, namely those where recursion can take place. For certain applications, such as free theorems, data refinement, quotients, and generalised rewriting, it is desirable that these operations do not ignore the other parameters. In this article, we formalise the generalisation BNFCC that extends the mapper and relator to covariant and contravariant parameters. We show that
  1. BNFCCs are closed under functor composition and least and greatest fixpoints,
  2. subtypes inherit the BNFCC structure under conditions that generalise those for the BNF case, and
  3. BNFCCs preserve quotients under mild conditions.
These proofs are carried out for abstract BNFCCs similar to the AFP entry BNF Operations. In addition, we apply the BNFCC theory to several concrete functors. [Modular_Assembly_Kit_Security] title = An Isabelle/HOL Formalization of the Modular Assembly Kit for Security Properties author = Oliver Bračevac , Richard Gay , Sylvia Grewe , Heiko Mantel , Henning Sudbrock , Markus Tasch topic = Computer science/Security date = 2018-05-07 notify = tasch@mais.informatik.tu-darmstadt.de abstract = The "Modular Assembly Kit for Security Properties" (MAKS) is a framework for both the definition and verification of possibilistic information-flow security properties at the specification-level. MAKS supports the uniform representation of a wide range of possibilistic information-flow properties and provides support for the verification of such properties via unwinding results and compositionality results. We provide a formalization of this framework in Isabelle/HOL. [AxiomaticCategoryTheory] title = Axiom Systems for Category Theory in Free Logic author = Christoph Benzmüller , Dana Scott topic = Mathematics/Category theory date = 2018-05-23 notify = c.benzmueller@gmail.com abstract = This document provides a concise overview on the core results of our previous work on the exploration of axioms systems for category theory. Extending the previous studies (http://arxiv.org/abs/1609.01493) we include one further axiomatic theory in our experiments. This additional theory has been suggested by Mac Lane in 1948. We show that the axioms proposed by Mac Lane are equivalent to the ones we studied before, which includes an axioms set suggested by Scott in the 1970s and another axioms set proposed by Freyd and Scedrov in 1990, which we slightly modified to remedy a minor technical issue. [OpSets] title = OpSets: Sequential Specifications for Replicated Datatypes author = Martin Kleppmann , Victor B. F. Gomes , Dominic P. Mulligan , Alastair R. Beresford topic = Computer science/Algorithms/Distributed, Computer science/Data structures date = 2018-05-10 notify = vb358@cam.ac.uk abstract = We introduce OpSets, an executable framework for specifying and reasoning about the semantics of replicated datatypes that provide eventual consistency in a distributed system, and for mechanically verifying algorithms that implement these datatypes. Our approach is simple but expressive, allowing us to succinctly specify a variety of abstract datatypes, including maps, sets, lists, text, graphs, trees, and registers. Our datatypes are also composable, enabling the construction of complex data structures. To demonstrate the utility of OpSets for analysing replication algorithms, we highlight an important correctness property for collaborative text editing that has traditionally been overlooked; algorithms that do not satisfy this property can exhibit awkward interleaving of text. We use OpSets to specify this correctness property and prove that although one existing replication algorithm satisfies this property, several other published algorithms do not. [Irrationality_J_Hancl] title = Irrational Rapidly Convergent Series author = Angeliki Koutsoukou-Argyraki , Wenda Li topic = Mathematics/Number theory, Mathematics/Analysis date = 2018-05-23 notify = ak2110@cam.ac.uk, wl302@cam.ac.uk abstract = We formalize with Isabelle/HOL a proof of a theorem by J. Hancl asserting the irrationality of the sum of a series consisting of rational numbers, built up by sequences that fulfill certain properties. Even though the criterion is a number theoretic result, the proof makes use only of analytical arguments. We also formalize a corollary of the theorem for a specific series fulfilling the assumptions of the theorem. [Optimal_BST] title = Optimal Binary Search Trees author = Tobias Nipkow , Dániel Somogyi <> topic = Computer science/Algorithms, Computer science/Data structures date = 2018-05-27 notify = nipkow@in.tum.de abstract = This article formalizes recursive algorithms for the construction of optimal binary search trees given fixed access frequencies. We follow Knuth (1971), Yao (1980) and Mehlhorn (1984). The algorithms are memoized with the help of the AFP article Monadification, Memoization and Dynamic Programming, thus yielding dynamic programming algorithms. [Projective_Geometry] title = Projective Geometry author = Anthony Bordg topic = Mathematics/Geometry date = 2018-06-14 notify = apdb3@cam.ac.uk abstract = We formalize the basics of projective geometry. In particular, we give a proof of the so-called Hessenberg's theorem in projective plane geometry. We also provide a proof of the so-called Desargues's theorem based on an axiomatization of (higher) projective space geometry using the notion of rank of a matroid. This last approach allows to handle incidence relations in an homogeneous way dealing only with points and without the need of talking explicitly about lines, planes or any higher entity. [Localization_Ring] title = The Localization of a Commutative Ring author = Anthony Bordg topic = Mathematics/Algebra date = 2018-06-14 notify = apdb3@cam.ac.uk abstract = We formalize the localization of a commutative ring R with respect to a multiplicative subset (i.e. a submonoid of R seen as a multiplicative monoid). This localization is itself a commutative ring and we build the natural homomorphism of rings from R to its localization. [Minsky_Machines] title = Minsky Machines author = Bertram Felgenhauer<> topic = Logic/Computability date = 2018-08-14 notify = int-e@gmx.de abstract =

We formalize undecidablity results for Minsky machines. To this end, we also formalize recursive inseparability.

We start by proving that Minsky machines can compute arbitrary primitive recursive and recursive functions. We then show that there is a deterministic Minsky machine with one argument and two final states such that the set of inputs that are accepted in one state is recursively inseparable from the set of inputs that are accepted in the other state.

As a corollary, the set of Minsky configurations that reach the first state but not the second recursively inseparable from the set of Minsky configurations that reach the second state but not the first. In particular both these sets are undecidable.

We do not prove that recursive functions can simulate Minsky machines.

[Neumann_Morgenstern_Utility] title = Von-Neumann-Morgenstern Utility Theorem author = Julian Parsert, Cezary Kaliszyk topic = Mathematics/Games and economics license = LGPL date = 2018-07-04 notify = julian.parsert@uibk.ac.at, cezary.kaliszyk@uibk.ac.at abstract = Utility functions form an essential part of game theory and economics. In order to guarantee the existence of utility functions most of the time sufficient properties are assumed in an axiomatic manner. One famous and very common set of such assumptions is that of expected utility theory. Here, the rationality, continuity, and independence of preferences is assumed. The von-Neumann-Morgenstern Utility theorem shows that these assumptions are necessary and sufficient for an expected utility function to exists. This theorem was proven by Neumann and Morgenstern in ``Theory of Games and Economic Behavior'' which is regarded as one of the most influential works in game theory. The formalization includes formal definitions of the underlying concepts including continuity and independence of preferences. [Simplex] title = An Incremental Simplex Algorithm with Unsatisfiable Core Generation author = Filip Marić , Mirko Spasić , René Thiemann topic = Computer science/Algorithms/Optimization date = 2018-08-24 notify = rene.thiemann@uibk.ac.at abstract = We present an Isabelle/HOL formalization and total correctness proof for the incremental version of the Simplex algorithm which is used in most state-of-the-art SMT solvers. It supports extraction of satisfying assignments, extraction of minimal unsatisfiable cores, incremental assertion of constraints and backtracking. The formalization relies on stepwise program refinement, starting from a simple specification, going through a number of refinement steps, and ending up in a fully executable functional implementation. Symmetries present in the algorithm are handled with special care. [Budan_Fourier] title = The Budan-Fourier Theorem and Counting Real Roots with Multiplicity author = Wenda Li topic = Mathematics/Analysis date = 2018-09-02 notify = wl302@cam.ac.uk, liwenda1990@hotmail.com abstract = This entry is mainly about counting and approximating real roots (of a polynomial) with multiplicity. We have first formalised the Budan-Fourier theorem: given a polynomial with real coefficients, we can calculate sign variations on Fourier sequences to over-approximate the number of real roots (counting multiplicity) within an interval. When all roots are known to be real, the over-approximation becomes tight: we can utilise this theorem to count real roots exactly. It is also worth noting that Descartes' rule of sign is a direct consequence of the Budan-Fourier theorem, and has been included in this entry. In addition, we have extended previous formalised Sturm's theorem to count real roots with multiplicity, while the original Sturm's theorem only counts distinct real roots. Compared to the Budan-Fourier theorem, our extended Sturm's theorem always counts roots exactly but may suffer from greater computational cost. [Quaternions] title = Quaternions author = Lawrence C. Paulson topic = Mathematics/Algebra, Mathematics/Geometry date = 2018-09-05 notify = lp15@cam.ac.uk abstract = This theory is inspired by the HOL Light development of quaternions, but follows its own route. Quaternions are developed coinductively, as in the existing formalisation of the complex numbers. Quaternions are quickly shown to belong to the type classes of real normed division algebras and real inner product spaces. And therefore they inherit a great body of facts involving algebraic laws, limits, continuity, etc., which must be proved explicitly in the HOL Light version. The development concludes with the geometric interpretation of the product of imaginary quaternions. [Octonions] title = Octonions author = Angeliki Koutsoukou-Argyraki topic = Mathematics/Algebra, Mathematics/Geometry date = 2018-09-14 notify = ak2110@cam.ac.uk abstract = We develop the basic theory of Octonions, including various identities and properties of the octonions and of the octonionic product, a description of 7D isometries and representations of orthogonal transformations. To this end we first develop the theory of the vector cross product in 7 dimensions. The development of the theory of Octonions is inspired by that of the theory of Quaternions by Lawrence Paulson. However, we do not work within the type class real_algebra_1 because the octonionic product is not associative. [Aggregation_Algebras] title = Aggregation Algebras author = Walter Guttmann topic = Mathematics/Algebra date = 2018-09-15 notify = walter.guttmann@canterbury.ac.nz abstract = We develop algebras for aggregation and minimisation for weight matrices and for edge weights in graphs. We verify the correctness of Prim's and Kruskal's minimum spanning tree algorithms based on these algebras. We also show numerous instances of these algebras based on linearly ordered commutative semigroups. extra-history = Change history: [2020-12-09]: moved Hoare logic to HOL-Hoare, moved spanning trees to Relational_Minimum_Spanning_Trees (revision dbb9bfaf4283) [Prime_Number_Theorem] title = The Prime Number Theorem author = Manuel Eberl , Lawrence C. Paulson topic = Mathematics/Number theory date = 2018-09-19 notify = manuel@pruvisto.org abstract =

This article provides a short proof of the Prime Number Theorem in several equivalent forms, most notably π(x) ~ x/ln x where π(x) is the number of primes no larger than x. It also defines other basic number-theoretic functions related to primes like Chebyshev's functions ϑ and ψ and the “n-th prime number” function pn. We also show various bounds and relationship between these functions are shown. Lastly, we derive Mertens' First and Second Theorem, i. e. ∑px ln p/p = ln x + O(1) and ∑px 1/p = ln ln x + M + O(1/ln x). We also give explicit bounds for the remainder terms.

The proof of the Prime Number Theorem builds on a library of Dirichlet series and analytic combinatorics. We essentially follow the presentation by Newman. The core part of the proof is a Tauberian theorem for Dirichlet series, which is proven using complex analysis and then used to strengthen Mertens' First Theorem to ∑px ln p/p = ln x + c + o(1).

A variant of this proof has been formalised before by Harrison in HOL Light, and formalisations of Selberg's elementary proof exist both by Avigad et al. in Isabelle and by Carneiro in Metamath. The advantage of the analytic proof is that, while it requires more powerful mathematical tools, it is considerably shorter and clearer. This article attempts to provide a short and clear formalisation of all components of that proof using the full range of mathematical machinery available in Isabelle, staying as close as possible to Newman's simple paper proof.

[Signature_Groebner] title = Signature-Based Gröbner Basis Algorithms author = Alexander Maletzky topic = Mathematics/Algebra, Computer science/Algorithms/Mathematical date = 2018-09-20 notify = alexander.maletzky@risc.jku.at abstract =

This article formalizes signature-based algorithms for computing Gröbner bases. Such algorithms are, in general, superior to other algorithms in terms of efficiency, and have not been formalized in any proof assistant so far. The present development is both generic, in the sense that most known variants of signature-based algorithms are covered by it, and effectively executable on concrete input thanks to Isabelle's code generator. Sample computations of benchmark problems show that the verified implementation of signature-based algorithms indeed outperforms the existing implementation of Buchberger's algorithm in Isabelle/HOL.

Besides total correctness of the algorithms, the article also proves that under certain conditions they a-priori detect and avoid all useless zero-reductions, and always return 'minimal' (in some sense) Gröbner bases if an input parameter is chosen in the right way.

The formalization follows the recent survey article by Eder and Faugère.

[Factored_Transition_System_Bounding] title = Upper Bounding Diameters of State Spaces of Factored Transition Systems author = Friedrich Kurz <>, Mohammad Abdulaziz topic = Computer science/Automata and formal languages, Mathematics/Graph theory date = 2018-10-12 notify = friedrich.kurz@tum.de, mohammad.abdulaziz@in.tum.de abstract = A completeness threshold is required to guarantee the completeness of planning as satisfiability, and bounded model checking of safety properties. One valid completeness threshold is the diameter of the underlying transition system. The diameter is the maximum element in the set of lengths of all shortest paths between pairs of states. The diameter is not calculated exactly in our setting, where the transition system is succinctly described using a (propositionally) factored representation. Rather, an upper bound on the diameter is calculated compositionally, by bounding the diameters of small abstract subsystems, and then composing those. We port a HOL4 formalisation of a compositional algorithm for computing a relatively tight upper bound on the system diameter. This compositional algorithm exploits acyclicity in the state space to achieve compositionality, and it was introduced by Abdulaziz et. al. The formalisation that we port is described as a part of another paper by Abdulaziz et. al. As a part of this porting we developed a libray about transition systems, which shall be of use in future related mechanisation efforts. [Smooth_Manifolds] title = Smooth Manifolds author = Fabian Immler , Bohua Zhan topic = Mathematics/Analysis, Mathematics/Topology date = 2018-10-22 notify = immler@in.tum.de, bzhan@ios.ac.cn abstract = We formalize the definition and basic properties of smooth manifolds in Isabelle/HOL. Concepts covered include partition of unity, tangent and cotangent spaces, and the fundamental theorem of path integrals. We also examine some concrete manifolds such as spheres and projective spaces. The formalization makes extensive use of the analysis and linear algebra libraries in Isabelle/HOL, in particular its “types-to-sets” mechanism. [Matroids] title = Matroids author = Jonas Keinholz<> topic = Mathematics/Combinatorics date = 2018-11-16 notify = manuel@pruvisto.org abstract =

This article defines the combinatorial structures known as Independence Systems and Matroids and provides basic concepts and theorems related to them. These structures play an important role in combinatorial optimisation, e. g. greedy algorithms such as Kruskal's algorithm. The development is based on Oxley's `What is a Matroid?'.

[Graph_Saturation] title = Graph Saturation author = Sebastiaan J. C. Joosten<> topic = Logic/Rewriting, Mathematics/Graph theory date = 2018-11-23 notify = sjcjoosten@gmail.com abstract = This is an Isabelle/HOL formalisation of graph saturation, closely following a paper by the author on graph saturation. Nine out of ten lemmas of the original paper are proven in this formalisation. The formalisation additionally includes two theorems that show the main premise of the paper: that consistency and entailment are decided through graph saturation. This formalisation does not give executable code, and it did not implement any of the optimisations suggested in the paper. [Functional_Ordered_Resolution_Prover] title = A Verified Functional Implementation of Bachmair and Ganzinger's Ordered Resolution Prover author = Anders Schlichtkrull , Jasmin Christian Blanchette , Dmitriy Traytel topic = Logic/General logic/Mechanization of proofs date = 2018-11-23 notify = andschl@dtu.dk,j.c.blanchette@vu.nl,traytel@inf.ethz.ch abstract = This Isabelle/HOL formalization refines the abstract ordered resolution prover presented in Section 4.3 of Bachmair and Ganzinger's "Resolution Theorem Proving" chapter in the Handbook of Automated Reasoning. The result is a functional implementation of a first-order prover. [Auto2_HOL] title = Auto2 Prover author = Bohua Zhan topic = Tools date = 2018-11-20 notify = bzhan@ios.ac.cn abstract = Auto2 is a saturation-based heuristic prover for higher-order logic, implemented as a tactic in Isabelle. This entry contains the instantiation of auto2 for Isabelle/HOL, along with two basic examples: solutions to some of the Pelletier’s problems, and elementary number theory of primes. [Order_Lattice_Props] title = Properties of Orderings and Lattices author = Georg Struth topic = Mathematics/Order date = 2018-12-11 notify = g.struth@sheffield.ac.uk abstract = These components add further fundamental order and lattice-theoretic concepts and properties to Isabelle's libraries. They follow by and large the introductory sections of the Compendium of Continuous Lattices, covering directed and filtered sets, down-closed and up-closed sets, ideals and filters, Galois connections, closure and co-closure operators. Some emphasis is on duality and morphisms between structures, as in the Compendium. To this end, three ad-hoc approaches to duality are compared. [Quantales] title = Quantales author = Georg Struth topic = Mathematics/Algebra date = 2018-12-11 notify = g.struth@sheffield.ac.uk abstract = These mathematical components formalise basic properties of quantales, together with some important models, constructions, and concepts, including quantic nuclei and conuclei. [Transformer_Semantics] title = Transformer Semantics author = Georg Struth topic = Mathematics/Algebra, Computer science/Semantics date = 2018-12-11 notify = g.struth@sheffield.ac.uk abstract = These mathematical components formalise predicate transformer semantics for programs, yet currently only for partial correctness and in the absence of faults. A first part for isotone (or monotone), Sup-preserving and Inf-preserving transformers follows Back and von Wright's approach, with additional emphasis on the quantalic structure of algebras of transformers. The second part develops Sup-preserving and Inf-preserving predicate transformers from the powerset monad, via its Kleisli category and Eilenberg-Moore algebras, with emphasis on adjunctions and dualities, as well as isomorphisms between relations, state transformers and predicate transformers. [Concurrent_Revisions] title = Formalization of Concurrent Revisions author = Roy Overbeek topic = Computer science/Concurrency date = 2018-12-25 notify = Roy.Overbeek@cwi.nl abstract = Concurrent revisions is a concurrency control model developed by Microsoft Research. It has many interesting properties that distinguish it from other well-known models such as transactional memory. One of these properties is determinacy: programs written within the model always produce the same outcome, independent of scheduling activity. The concurrent revisions model has an operational semantics, with an informal proof of determinacy. This document contains an Isabelle/HOL formalization of this semantics and the proof of determinacy. [Core_DOM] title = A Formal Model of the Document Object Model author = Achim D. Brucker , Michael Herzberg topic = Computer science/Data structures date = 2018-12-26 notify = adbrucker@0x5f.org abstract = In this AFP entry, we formalize the core of the Document Object Model (DOM). At its core, the DOM defines a tree-like data structure for representing documents in general and HTML documents in particular. It is the heart of any modern web browser. Formalizing the key concepts of the DOM is a prerequisite for the formal reasoning over client-side JavaScript programs and for the analysis of security concepts in modern web browsers. We present a formalization of the core DOM, with focus on the node-tree and the operations defined on node-trees, in Isabelle/HOL. We use the formalization to verify the functional correctness of the most important functions defined in the DOM standard. Moreover, our formalization is 1) extensible, i.e., can be extended without the need of re-proving already proven properties and 2) executable, i.e., we can generate executable code from our specification. [Core_SC_DOM] title = The Safely Composable DOM author = Achim D. Brucker , Michael Herzberg topic = Computer science/Data structures date = 2020-09-28 notify = adbrucker@0x5f.org, mail@michael-herzberg.de abstract = In this AFP entry, we formalize the core of the Safely Composable Document Object Model (SC DOM). The SC DOM improve the standard DOM (as formalized in the AFP entry "Core DOM") by strengthening the tree boundaries set by shadow roots: in the SC DOM, the shadow root is a sub-class of the document class (instead of a base class). This modifications also results in changes to some API methods (e.g., getOwnerDocument) to return the nearest shadow root rather than the document root. As a result, many API methods that, when called on a node inside a shadow tree, would previously ``break out'' and return or modify nodes that are possibly outside the shadow tree, now stay within its boundaries. This change in behavior makes programs that operate on shadow trees more predictable for the developer and allows them to make more assumptions about other code accessing the DOM. [Shadow_SC_DOM] title = A Formal Model of the Safely Composable Document Object Model with Shadow Roots author = Achim D. Brucker , Michael Herzberg topic = Computer science/Data structures date = 2020-09-28 notify = adbrucker@0x5f.org, mail@michael-herzberg.de abstract = In this AFP entry, we extend our formalization of the safely composable DOM with Shadow Roots. This is a proposal for Shadow Roots with stricter safety guarantess than the standard compliant formalization (see "Shadow DOM"). Shadow Roots are a recent proposal of the web community to support a component-based development approach for client-side web applications. Shadow roots are a significant extension to the DOM standard and, as web standards are condemned to be backward compatible, such extensions often result in complex specification that may contain unwanted subtleties that can be detected by a formalization. Our Isabelle/HOL formalization is, in the sense of object-orientation, an extension of our formalization of the core DOM and enjoys the same basic properties, i.e., it is extensible, i.e., can be extended without the need of re-proving already proven properties and executable, i.e., we can generate executable code from our specification. We exploit the executability to show that our formalization complies to the official standard of the W3C, respectively, the WHATWG. [SC_DOM_Components] title = A Formalization of Safely Composable Web Components author = Achim D. Brucker , Michael Herzberg topic = Computer science/Data structures date = 2020-09-28 notify = adbrucker@0x5f.org, mail@michael-herzberg.de abstract = While the (safely composable) DOM with shadow trees provide the technical basis for defining web components, it does neither defines the concept of web components nor specifies the safety properties that web components should guarantee. Consequently, the standard also does not discuss how or even if the methods for modifying the DOM respect component boundaries. In AFP entry, we present a formally verified model of safely composable web components and define safety properties which ensure that different web components can only interact with each other using well-defined interfaces. Moreover, our verification of the application programming interface (API) of the DOM revealed numerous invariants that implementations of the DOM API need to preserve to ensure the integrity of components. In comparison to the strict standard compliance formalization of Web Components in the AFP entry "DOM_Components", the notion of components in this entry (based on "SC_DOM" and "Shadow_SC_DOM") provides much stronger safety guarantees. [Store_Buffer_Reduction] title = A Reduction Theorem for Store Buffers author = Ernie Cohen , Norbert Schirmer topic = Computer science/Concurrency date = 2019-01-07 notify = norbert.schirmer@web.de abstract = When verifying a concurrent program, it is usual to assume that memory is sequentially consistent. However, most modern multiprocessors depend on store buffering for efficiency, and provide native sequential consistency only at a substantial performance penalty. To regain sequential consistency, a programmer has to follow an appropriate programming discipline. However, naïve disciplines, such as protecting all shared accesses with locks, are not flexible enough for building high-performance multiprocessor software. We present a new discipline for concurrent programming under TSO (total store order, with store buffer forwarding). It does not depend on concurrency primitives, such as locks. Instead, threads use ghost operations to acquire and release ownership of memory addresses. A thread can write to an address only if no other thread owns it, and can read from an address only if it owns it or it is shared and the thread has flushed its store buffer since it last wrote to an address it did not own. This discipline covers both coarse-grained concurrency (where data is protected by locks) as well as fine-grained concurrency (where atomic operations race to memory). We formalize this discipline in Isabelle/HOL, and prove that if every execution of a program in a system without store buffers follows the discipline, then every execution of the program with store buffers is sequentially consistent. Thus, we can show sequential consistency under TSO by ordinary assertional reasoning about the program, without having to consider store buffers at all. [IMP2] title = IMP2 – Simple Program Verification in Isabelle/HOL author = Peter Lammich , Simon Wimmer topic = Computer science/Programming languages/Logics, Computer science/Algorithms date = 2019-01-15 notify = lammich@in.tum.de abstract = IMP2 is a simple imperative language together with Isabelle tooling to create a program verification environment in Isabelle/HOL. The tools include a C-like syntax, a verification condition generator, and Isabelle commands for the specification of programs. The framework is modular, i.e., it allows easy reuse of already proved programs within larger programs. This entry comes with a quickstart guide and a large collection of examples, spanning basic algorithms with simple proofs to more advanced algorithms and proof techniques like data refinement. Some highlights from the examples are:
  • Bisection Square Root,
  • Extended Euclid,
  • Exponentiation by Squaring,
  • Binary Search,
  • Insertion Sort,
  • Quicksort,
  • Depth First Search.
The abstract syntax and semantics are very simple and well-documented. They are suitable to be used in a course, as extension to the IMP language which comes with the Isabelle distribution. While this entry is limited to a simple imperative language, the ideas could be extended to more sophisticated languages. [Farkas] title = Farkas' Lemma and Motzkin's Transposition Theorem author = Ralph Bottesch , Max W. Haslbeck , René Thiemann topic = Mathematics/Algebra date = 2019-01-17 notify = rene.thiemann@uibk.ac.at abstract = We formalize a proof of Motzkin's transposition theorem and Farkas' lemma in Isabelle/HOL. Our proof is based on the formalization of the simplex algorithm which, given a set of linear constraints, either returns a satisfying assignment to the problem or detects unsatisfiability. By reusing facts about the simplex algorithm we show that a set of linear constraints is unsatisfiable if and only if there is a linear combination of the constraints which evaluates to a trivially unsatisfiable inequality. [Auto2_Imperative_HOL] title = Verifying Imperative Programs using Auto2 author = Bohua Zhan topic = Computer science/Algorithms, Computer science/Data structures date = 2018-12-21 notify = bzhan@ios.ac.cn abstract = This entry contains the application of auto2 to verifying functional and imperative programs. Algorithms and data structures that are verified include linked lists, binary search trees, red-black trees, interval trees, priority queue, quicksort, union-find, Dijkstra's algorithm, and a sweep-line algorithm for detecting rectangle intersection. The imperative verification is based on Imperative HOL and its separation logic framework. A major goal of this work is to set up automation in order to reduce the length of proof that the user needs to provide, both for verifying functional programs and for working with separation logic. [UTP] title = Isabelle/UTP: Mechanised Theory Engineering for Unifying Theories of Programming author = Simon Foster , Frank Zeyda<>, Yakoub Nemouchi , Pedro Ribeiro<>, Burkhart Wolff topic = Computer science/Programming languages/Logics date = 2019-02-01 notify = simon.foster@york.ac.uk abstract = Isabelle/UTP is a mechanised theory engineering toolkit based on Hoare and He’s Unifying Theories of Programming (UTP). UTP enables the creation of denotational, algebraic, and operational semantics for different programming languages using an alphabetised relational calculus. We provide a semantic embedding of the alphabetised relational calculus in Isabelle/HOL, including new type definitions, relational constructors, automated proof tactics, and accompanying algebraic laws. Isabelle/UTP can be used to both capture laws of programming for different languages, and put these fundamental theorems to work in the creation of associated verification tools, using calculi like Hoare logics. This document describes the relational core of the UTP in Isabelle/HOL. [HOL-CSP] title = HOL-CSP Version 2.0 author = Safouan Taha , Lina Ye , Burkhart Wolff topic = Computer science/Concurrency/Process calculi, Computer science/Semantics date = 2019-04-26 notify = wolff@lri.fr abstract = This is a complete formalization of the work of Hoare and Roscoe on the denotational semantics of the Failure/Divergence Model of CSP. It follows essentially the presentation of CSP in Roscoe’s Book ”Theory and Practice of Concurrency” [8] and the semantic details in a joint Paper of Roscoe and Brooks ”An improved failures model for communicating processes". The present work is based on a prior formalization attempt, called HOL-CSP 1.0, done in 1997 by H. Tej and B. Wolff with the Isabelle proof technology available at that time. This work revealed minor, but omnipresent foundational errors in key concepts like the process invariant. The present version HOL-CSP profits from substantially improved libraries (notably HOLCF), improved automated proof techniques, and structured proof techniques in Isar and is substantially shorter but more complete. [Probabilistic_Prime_Tests] title = Probabilistic Primality Testing author = Daniel Stüwe<>, Manuel Eberl topic = Mathematics/Number theory date = 2019-02-11 notify = manuel@pruvisto.org abstract =

The most efficient known primality tests are probabilistic in the sense that they use randomness and may, with some probability, mistakenly classify a composite number as prime – but never a prime number as composite. Examples of this are the Miller–Rabin test, the Solovay–Strassen test, and (in most cases) Fermat's test.

This entry defines these three tests and proves their correctness. It also develops some of the number-theoretic foundations, such as Carmichael numbers and the Jacobi symbol with an efficient executable algorithm to compute it.

[Kruskal] title = Kruskal's Algorithm for Minimum Spanning Forest author = Maximilian P.L. Haslbeck , Peter Lammich , Julian Biendarra<> topic = Computer science/Algorithms/Graph date = 2019-02-14 notify = haslbema@in.tum.de, lammich@in.tum.de abstract = This Isabelle/HOL formalization defines a greedy algorithm for finding a minimum weight basis on a weighted matroid and proves its correctness. This algorithm is an abstract version of Kruskal's algorithm. We interpret the abstract algorithm for the cycle matroid (i.e. forests in a graph) and refine it to imperative executable code using an efficient union-find data structure. Our formalization can be instantiated for different graph representations. We provide instantiations for undirected graphs and symmetric directed graphs. [List_Inversions] title = The Inversions of a List author = Manuel Eberl topic = Computer science/Algorithms date = 2019-02-01 notify = manuel@pruvisto.org abstract =

This entry defines the set of inversions of a list, i.e. the pairs of indices that violate sortedness. It also proves the correctness of the well-known O(n log n) divide-and-conquer algorithm to compute the number of inversions.

[Prime_Distribution_Elementary] title = Elementary Facts About the Distribution of Primes author = Manuel Eberl topic = Mathematics/Number theory date = 2019-02-21 notify = manuel@pruvisto.org abstract =

This entry is a formalisation of Chapter 4 (and parts of Chapter 3) of Apostol's Introduction to Analytic Number Theory. The main topics that are addressed are properties of the distribution of prime numbers that can be shown in an elementary way (i. e. without the Prime Number Theorem), the various equivalent forms of the PNT (which imply each other in elementary ways), and consequences that follow from the PNT in elementary ways. The latter include, most notably, asymptotic bounds for the number of distinct prime factors of n, the divisor function d(n), Euler's totient function φ(n), and lcm(1,…,n).

[Safe_OCL] title = Safe OCL author = Denis Nikiforov <> topic = Computer science/Programming languages/Language definitions license = LGPL date = 2019-03-09 notify = denis.nikif@gmail.com abstract =

The theory is a formalization of the OCL type system, its abstract syntax and expression typing rules. The theory does not define a concrete syntax and a semantics. In contrast to Featherweight OCL, it is based on a deep embedding approach. The type system is defined from scratch, it is not based on the Isabelle HOL type system.

The Safe OCL distincts nullable and non-nullable types. Also the theory gives a formal definition of safe navigation operations. The Safe OCL typing rules are much stricter than rules given in the OCL specification. It allows one to catch more errors on a type checking phase.

The type theory presented is four-layered: classes, basic types, generic types, errorable types. We introduce the following new types: non-nullable types (T[1]), nullable types (T[?]), OclSuper. OclSuper is a supertype of all other types (basic types, collections, tuples). This type allows us to define a total supremum function, so types form an upper semilattice. It allows us to define rich expression typing rules in an elegant manner.

The Preliminaries Chapter of the theory defines a number of helper lemmas for transitive closures and tuples. It defines also a generic object model independent from OCL. It allows one to use the theory as a reference for formalization of analogous languages.

[QHLProver] title = Quantum Hoare Logic author = Junyi Liu<>, Bohua Zhan , Shuling Wang<>, Shenggang Ying<>, Tao Liu<>, Yangjia Li<>, Mingsheng Ying<>, Naijun Zhan<> topic = Computer science/Programming languages/Logics, Computer science/Semantics date = 2019-03-24 notify = bzhan@ios.ac.cn abstract = We formalize quantum Hoare logic as given in [1]. In particular, we specify the syntax and denotational semantics of a simple model of quantum programs. Then, we write down the rules of quantum Hoare logic for partial correctness, and show the soundness and completeness of the resulting proof system. As an application, we verify the correctness of Grover’s algorithm. [Transcendence_Series_Hancl_Rucki] title = The Transcendence of Certain Infinite Series author = Angeliki Koutsoukou-Argyraki , Wenda Li topic = Mathematics/Analysis, Mathematics/Number theory date = 2019-03-27 notify = wl302@cam.ac.uk, ak2110@cam.ac.uk abstract = We formalize the proofs of two transcendence criteria by J. Hančl and P. Rucki that assert the transcendence of the sums of certain infinite series built up by sequences that fulfil certain properties. Both proofs make use of Roth's celebrated theorem on diophantine approximations to algebraic numbers from 1955 which we implement as an assumption without having formalised its proof. [Binding_Syntax_Theory] title = A General Theory of Syntax with Bindings author = Lorenzo Gheri , Andrei Popescu topic = Computer science/Programming languages/Lambda calculi, Computer science/Functional programming, Logic/General logic/Mechanization of proofs date = 2019-04-06 notify = a.popescu@mdx.ac.uk, lor.gheri@gmail.com abstract = We formalize a theory of syntax with bindings that has been developed and refined over the last decade to support several large formalization efforts. Terms are defined for an arbitrary number of constructors of varying numbers of inputs, quotiented to alpha-equivalence and sorted according to a binding signature. The theory includes many properties of the standard operators on terms: substitution, swapping and freshness. It also includes bindings-aware induction and recursion principles and support for semantic interpretation. This work has been presented in the ITP 2017 paper “A Formalized General Theory of Syntax with Bindings”. [LTL_Master_Theorem] title = A Compositional and Unified Translation of LTL into ω-Automata author = Benedikt Seidl , Salomon Sickert topic = Computer science/Automata and formal languages date = 2019-04-16 notify = benedikt.seidl@tum.de, s.sickert@tum.de abstract = We present a formalisation of the unified translation approach of linear temporal logic (LTL) into ω-automata from [1]. This approach decomposes LTL formulas into ``simple'' languages and allows a clear separation of concerns: first, we formalise the purely logical result yielding this decomposition; second, we instantiate this generic theory to obtain a construction for deterministic (state-based) Rabin automata (DRA). We extract from this particular instantiation an executable tool translating LTL to DRAs. To the best of our knowledge this is the first verified translation from LTL to DRAs that is proven to be double exponential in the worst case which asymptotically matches the known lower bound.

[1] Javier Esparza, Jan Kretínský, Salomon Sickert. One Theorem to Rule Them All: A Unified Translation of LTL into ω-Automata. LICS 2018 [LambdaAuth] title = Formalization of Generic Authenticated Data Structures author = Matthias Brun<>, Dmitriy Traytel topic = Computer science/Security, Computer science/Programming languages/Lambda calculi date = 2019-05-14 notify = traytel@inf.ethz.ch abstract = Authenticated data structures are a technique for outsourcing data storage and maintenance to an untrusted server. The server is required to produce an efficiently checkable and cryptographically secure proof that it carried out precisely the requested computation. Miller et al. introduced λ• (pronounced lambda auth)—a functional programming language with a built-in primitive authentication construct, which supports a wide range of user-specified authenticated data structures while guaranteeing certain correctness and security properties for all well-typed programs. We formalize λ• and prove its correctness and security properties. With Isabelle's help, we uncover and repair several mistakes in the informal proofs and lemma statements. Our findings are summarized in an ITP'19 paper. [IMP2_Binary_Heap] title = Binary Heaps for IMP2 author = Simon Griebel<> topic = Computer science/Data structures, Computer science/Algorithms date = 2019-06-13 notify = s.griebel@tum.de abstract = In this submission array-based binary minimum heaps are formalized. The correctness of the following heap operations is proved: insert, get-min, delete-min and make-heap. These are then used to verify an in-place heapsort. The formalization is based on IMP2, an imperative program verification framework implemented in Isabelle/HOL. The verified heap functions are iterative versions of the partly recursive functions found in "Algorithms and Data Structures – The Basic Toolbox" by K. Mehlhorn and P. Sanders and "Introduction to Algorithms" by T. H. Cormen, C. E. Leiserson, R. L. Rivest and C. Stein. [Groebner_Macaulay] title = Gröbner Bases, Macaulay Matrices and Dubé's Degree Bounds author = Alexander Maletzky topic = Mathematics/Algebra date = 2019-06-15 notify = alexander.maletzky@risc.jku.at abstract = This entry formalizes the connection between Gröbner bases and Macaulay matrices (sometimes also referred to as `generalized Sylvester matrices'). In particular, it contains a method for computing Gröbner bases, which proceeds by first constructing some Macaulay matrix of the initial set of polynomials, then row-reducing this matrix, and finally converting the result back into a set of polynomials. The output is shown to be a Gröbner basis if the Macaulay matrix constructed in the first step is sufficiently large. In order to obtain concrete upper bounds on the size of the matrix (and hence turn the method into an effectively executable algorithm), Dubé's degree bounds on Gröbner bases are utilized; consequently, they are also part of the formalization. [Linear_Inequalities] title = Linear Inequalities author = Ralph Bottesch , Alban Reynaud <>, René Thiemann topic = Mathematics/Algebra date = 2019-06-21 notify = rene.thiemann@uibk.ac.at abstract = We formalize results about linear inqualities, mainly from Schrijver's book. The main results are the proof of the fundamental theorem on linear inequalities, Farkas' lemma, Carathéodory's theorem, the Farkas-Minkowsky-Weyl theorem, the decomposition theorem of polyhedra, and Meyer's result that the integer hull of a polyhedron is a polyhedron itself. Several theorems include bounds on the appearing numbers, and in particular we provide an a-priori bound on mixed-integer solutions of linear inequalities. [Linear_Programming] title = Linear Programming author = Julian Parsert , Cezary Kaliszyk topic = Mathematics/Algebra date = 2019-08-06 notify = julian.parsert@gmail.com, cezary.kaliszyk@uibk.ac.at abstract = We use the previous formalization of the general simplex algorithm to formulate an algorithm for solving linear programs. We encode the linear programs using only linear constraints. Solving these constraints also solves the original linear program. This algorithm is proven to be sound by applying the weak duality theorem which is also part of this formalization. [Differential_Game_Logic] title = Differential Game Logic author = André Platzer topic = Computer science/Programming languages/Logics date = 2019-06-03 notify = aplatzer@cs.cmu.edu abstract = This formalization provides differential game logic (dGL), a logic for proving properties of hybrid game. In addition to the syntax and semantics, it formalizes a uniform substitution calculus for dGL. Church's uniform substitutions substitute a term or formula for a function or predicate symbol everywhere. The uniform substitutions for dGL also substitute hybrid games for a game symbol everywhere. We prove soundness of one-pass uniform substitutions and the axioms of differential game logic with respect to their denotational semantics. One-pass uniform substitutions are faster by postponing soundness-critical admissibility checks with a linear pass homomorphic application and regain soundness by a variable condition at the replacements. The formalization is based on prior non-mechanized soundness proofs for dGL. [BenOr_Kozen_Reif] title = The BKR Decision Procedure for Univariate Real Arithmetic author = Katherine Cordwell , Yong Kiam Tan , André Platzer topic = Computer science/Algorithms/Mathematical date = 2021-04-24 notify = kcordwel@cs.cmu.edu, yongkiat@cs.cmu.edu, aplatzer@cs.cmu.edu abstract = We formalize the univariate case of Ben-Or, Kozen, and Reif's decision procedure for first-order real arithmetic (the BKR algorithm). We also formalize the univariate case of Renegar's variation of the BKR algorithm. The two formalizations differ mathematically in minor ways (that have significant impact on the multivariate case), but are quite similar in proof structure. Both rely on sign-determination (finding the set of consistent sign assignments for a set of polynomials). The method used for sign-determination is similar to Tarski's original quantifier elimination algorithm (it stores key information in a matrix equation), but with a reduction step to keep complexity low. [Complete_Non_Orders] title = Complete Non-Orders and Fixed Points author = Akihisa Yamada , Jérémy Dubut topic = Mathematics/Order date = 2019-06-27 notify = akihisayamada@nii.ac.jp, dubut@nii.ac.jp abstract = We develop an Isabelle/HOL library of order-theoretic concepts, such as various completeness conditions and fixed-point theorems. We keep our formalization as general as possible: we reprove several well-known results about complete orders, often without any properties of ordering, thus complete non-orders. In particular, we generalize the Knaster–Tarski theorem so that we ensure the existence of a quasi-fixed point of monotone maps over complete non-orders, and show that the set of quasi-fixed points is complete under a mild condition—attractivity—which is implied by either antisymmetry or transitivity. This result generalizes and strengthens a result by Stauti and Maaden. Finally, we recover Kleene’s fixed-point theorem for omega-complete non-orders, again using attractivity to prove that Kleene’s fixed points are least quasi-fixed points. [Priority_Search_Trees] title = Priority Search Trees author = Peter Lammich , Tobias Nipkow topic = Computer science/Data structures date = 2019-06-25 notify = lammich@in.tum.de abstract = We present a new, purely functional, simple and efficient data structure combining a search tree and a priority queue, which we call a priority search tree. The salient feature of priority search trees is that they offer a decrease-key operation, something that is missing from other simple, purely functional priority queue implementations. Priority search trees can be implemented on top of any search tree. This entry does the implementation for red-black trees. This entry formalizes the first part of our ITP-2019 proof pearl Purely Functional, Simple and Efficient Priority Search Trees and Applications to Prim and Dijkstra. [Prim_Dijkstra_Simple] title = Purely Functional, Simple, and Efficient Implementation of Prim and Dijkstra author = Peter Lammich , Tobias Nipkow topic = Computer science/Algorithms/Graph date = 2019-06-25 notify = lammich@in.tum.de abstract = We verify purely functional, simple and efficient implementations of Prim's and Dijkstra's algorithms. This constitutes the first verification of an executable and even efficient version of Prim's algorithm. This entry formalizes the second part of our ITP-2019 proof pearl Purely Functional, Simple and Efficient Priority Search Trees and Applications to Prim and Dijkstra. [MFOTL_Monitor] title = Formalization of a Monitoring Algorithm for Metric First-Order Temporal Logic author = Joshua Schneider , Dmitriy Traytel topic = Computer science/Algorithms, Logic/General logic/Temporal logic, Computer science/Automata and formal languages date = 2019-07-04 notify = joshua.schneider@inf.ethz.ch, traytel@inf.ethz.ch abstract = A monitor is a runtime verification tool that solves the following problem: Given a stream of time-stamped events and a policy formulated in a specification language, decide whether the policy is satisfied at every point in the stream. We verify the correctness of an executable monitor for specifications given as formulas in metric first-order temporal logic (MFOTL), an expressive extension of linear temporal logic with real-time constraints and first-order quantification. The verified monitor implements a simplified variant of the algorithm used in the efficient MonPoly monitoring tool. The formalization is presented in a RV 2019 paper, which also compares the output of the verified monitor to that of other monitoring tools on randomly generated inputs. This case study revealed several errors in the optimized but unverified tools. extra-history = Change history: [2020-08-13]: added the formalization of the abstract slicing framework and joint data slicer (revision b1639ed541b7)
[FOL_Seq_Calc1] title = A Sequent Calculus for First-Order Logic author = Asta Halkjær From contributors = Alexander Birch Jensen , Anders Schlichtkrull , Jørgen Villadsen topic = Logic/Proof theory date = 2019-07-18 notify = ahfrom@dtu.dk abstract = This work formalizes soundness and completeness of a one-sided sequent calculus for first-order logic. The completeness is shown via a translation from a complete semantic tableau calculus, the proof of which is based on the First-Order Logic According to Fitting theory. The calculi and proof techniques are taken from Ben-Ari's Mathematical Logic for Computer Science. [Szpilrajn] title = Order Extension and Szpilrajn's Extension Theorem author = Peter Zeller , Lukas Stevens topic = Mathematics/Order date = 2019-07-27 notify = p_zeller@cs.uni-kl.de abstract = This entry is concerned with the principle of order extension, i.e. the extension of an order relation to a total order relation. To this end, we prove a more general version of Szpilrajn's extension theorem employing terminology from the book "Consistency, Choice, and Rationality" by Bossert and Suzumura. We also formalize theorem 2.7 of their book. extra-history = Change history: [2021-03-22]: (by Lukas Stevens) generalise Szpilrajn's extension theorem and add material from the book "Consistency, Choice, and Rationality" [TESL_Language] title = A Formal Development of a Polychronous Polytimed Coordination Language author = Hai Nguyen Van , Frédéric Boulanger , Burkhart Wolff topic = Computer science/System description languages, Computer science/Semantics, Computer science/Concurrency date = 2019-07-30 notify = frederic.boulanger@centralesupelec.fr, burkhart.wolff@lri.fr abstract = The design of complex systems involves different formalisms for modeling their different parts or aspects. The global model of a system may therefore consist of a coordination of concurrent sub-models that use different paradigms. We develop here a theory for a language used to specify the timed coordination of such heterogeneous subsystems by addressing the following issues:

  • the behavior of the sub-systems is observed only at a series of discrete instants,
  • events may occur in different sub-systems at unrelated times, leading to polychronous systems, which do not necessarily have a common base clock,
  • coordination between subsystems involves causality, so the occurrence of an event may enforce the occurrence of other events, possibly after a certain duration has elapsed or an event has occurred a given number of times,
  • the domain of time (discrete, rational, continuous...) may be different in the subsystems, leading to polytimed systems,
  • the time frames of different sub-systems may be related (for instance, time in a GPS satellite and in a GPS receiver on Earth are related although they are not the same).
Firstly, a denotational semantics of the language is defined. Then, in order to be able to incrementally check the behavior of systems, an operational semantics is given, with proofs of progress, soundness and completeness with regard to the denotational semantics. These proofs are made according to a setup that can scale up when new operators are added to the language. In order for specifications to be composed in a clean way, the language should be invariant by stuttering (i.e., adding observation instants at which nothing happens). The proof of this invariance is also given. [Stellar_Quorums] title = Stellar Quorum Systems author = Giuliano Losa topic = Computer science/Algorithms/Distributed date = 2019-08-01 notify = giuliano@galois.com abstract = We formalize the static properties of personal Byzantine quorum systems (PBQSs) and Stellar quorum systems, as described in the paper ``Stellar Consensus by Reduction'' (to appear at DISC 2019). [IMO2019] title = Selected Problems from the International Mathematical Olympiad 2019 author = Manuel Eberl topic = Mathematics/Misc date = 2019-08-05 notify = manuel@pruvisto.org abstract =

This entry contains formalisations of the answers to three of the six problem of the International Mathematical Olympiad 2019, namely Q1, Q4, and Q5.

The reason why these problems were chosen is that they are particularly amenable to formalisation: they can be solved with minimal use of libraries. The remaining three concern geometry and graph theory, which, in the author's opinion, are more difficult to formalise resp. require a more complex library.

[Adaptive_State_Counting] title = Formalisation of an Adaptive State Counting Algorithm author = Robert Sachtleben topic = Computer science/Automata and formal languages, Computer science/Algorithms date = 2019-08-16 notify = rob_sac@uni-bremen.de abstract = This entry provides a formalisation of a refinement of an adaptive state counting algorithm, used to test for reduction between finite state machines. The algorithm has been originally presented by Hierons in the paper Testing from a Non-Deterministic Finite State Machine Using Adaptive State Counting. Definitions for finite state machines and adaptive test cases are given and many useful theorems are derived from these. The algorithm is formalised using mutually recursive functions, for which it is proven that the generated test suite is sufficient to test for reduction against finite state machines of a certain fault domain. Additionally, the algorithm is specified in a simple WHILE-language and its correctness is shown using Hoare-logic. [Jacobson_Basic_Algebra] title = A Case Study in Basic Algebra author = Clemens Ballarin topic = Mathematics/Algebra date = 2019-08-30 notify = ballarin@in.tum.de abstract = The focus of this case study is re-use in abstract algebra. It contains locale-based formalisations of selected parts of set, group and ring theory from Jacobson's Basic Algebra leading to the respective fundamental homomorphism theorems. The study is not intended as a library base for abstract algebra. It rather explores an approach towards abstract algebra in Isabelle. [Hybrid_Systems_VCs] title = Verification Components for Hybrid Systems author = Jonathan Julian Huerta y Munive <> topic = Mathematics/Algebra, Mathematics/Analysis date = 2019-09-10 notify = jjhuertaymunive1@sheffield.ac.uk, jonjulian23@gmail.com abstract = These components formalise a semantic framework for the deductive verification of hybrid systems. They support reasoning about continuous evolutions of hybrid programs in the style of differential dynamics logic. Vector fields or flows model these evolutions, and their verification is done with invariants for the former or orbits for the latter. Laws of modal Kleene algebra or categorical predicate transformers implement the verification condition generation. Examples show the approach at work. extra-history = Change history: [2020-12-13]: added components based on Kleene algebras with tests. These implement differential Hoare logic (dH) and a Morgan-style differential refinement calculus (dR) for verification of hybrid programs. [Generic_Join] title = Formalization of Multiway-Join Algorithms author = Thibault Dardinier<> topic = Computer science/Algorithms date = 2019-09-16 notify = tdardini@student.ethz.ch, traytel@inf.ethz.ch abstract = Worst-case optimal multiway-join algorithms are recent seminal achievement of the database community. These algorithms compute the natural join of multiple relational databases and improve in the worst case over traditional query plan optimizations of nested binary joins. In 2014, Ngo, Ré, and Rudra gave a unified presentation of different multi-way join algorithms. We formalized and proved correct their "Generic Join" algorithm and extended it to support negative joins. [Aristotles_Assertoric_Syllogistic] title = Aristotle's Assertoric Syllogistic author = Angeliki Koutsoukou-Argyraki topic = Logic/Philosophical aspects date = 2019-10-08 notify = ak2110@cam.ac.uk abstract = We formalise with Isabelle/HOL some basic elements of Aristotle's assertoric syllogistic following the article from the Stanford Encyclopedia of Philosophy by Robin Smith. To this end, we use a set theoretic formulation (covering both individual and general predication). In particular, we formalise the deductions in the Figures and after that we present Aristotle's metatheoretical observation that all deductions in the Figures can in fact be reduced to either Barbara or Celarent. As the formal proofs prove to be straightforward, the interest of this entry lies in illustrating the functionality of Isabelle and high efficiency of Sledgehammer for simple exercises in philosophy. [VerifyThis2019] title = VerifyThis 2019 -- Polished Isabelle Solutions author = Peter Lammich<>, Simon Wimmer topic = Computer science/Algorithms date = 2019-10-16 notify = lammich@in.tum.de, wimmers@in.tum.de abstract = VerifyThis 2019 (http://www.pm.inf.ethz.ch/research/verifythis.html) was a program verification competition associated with ETAPS 2019. It was the 8th event in the VerifyThis competition series. In this entry, we present polished and completed versions of our solutions that we created during the competition. [ZFC_in_HOL] title = Zermelo Fraenkel Set Theory in Higher-Order Logic author = Lawrence C. Paulson topic = Logic/Set theory date = 2019-10-24 notify = lp15@cam.ac.uk abstract =

This entry is a new formalisation of ZFC set theory in Isabelle/HOL. It is logically equivalent to Obua's HOLZF; the point is to have the closest possible integration with the rest of Isabelle/HOL, minimising the amount of new notations and exploiting type classes.

There is a type V of sets and a function elts :: V => V set mapping a set to its elements. Classes simply have type V set, and a predicate identifies the small classes: those that correspond to actual sets. Type classes connected with orders and lattices are used to minimise the amount of new notation for concepts such as the subset relation, union and intersection. Basic concepts — Cartesian products, disjoint sums, natural numbers, functions, etc. — are formalised.

More advanced set-theoretic concepts, such as transfinite induction, ordinals, cardinals and the transitive closure of a set, are also provided. The definition of addition and multiplication for general sets (not just ordinals) follows Kirby.

The theory provides two type classes with the aim of facilitating developments that combine V with other Isabelle/HOL types: embeddable, the class of types that can be injected into V (including V itself as well as V*V, etc.), and small, the class of types that correspond to some ZF set.

extra-history = Change history: [2020-01-28]: Generalisation of the "small" predicate and order types to arbitrary sets; ordinal exponentiation; introduction of the coercion ord_of_nat :: "nat => V"; numerous new lemmas. (revision 6081d5be8d08) [Interval_Arithmetic_Word32] title = Interval Arithmetic on 32-bit Words author = Brandon Bohrer topic = Computer science/Data structures date = 2019-11-27 notify = bjbohrer@gmail.com, bbohrer@cs.cmu.edu abstract = Interval_Arithmetic implements conservative interval arithmetic computations, then uses this interval arithmetic to implement a simple programming language where all terms have 32-bit signed word values, with explicit infinities for terms outside the representable bounds. Our target use case is interpreters for languages that must have a well-understood low-level behavior. We include a formalization of bounded-length strings which are used for the identifiers of our language. Bounded-length identifiers are useful in some applications, for example the Differential_Dynamic_Logic article, where a Euclidean space indexed by identifiers demands that identifiers are finitely many. [Generalized_Counting_Sort] title = An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges author = Pasquale Noce topic = Computer science/Algorithms, Computer science/Functional programming date = 2019-12-04 notify = pasquale.noce.lavoro@gmail.com abstract = Counting sort is a well-known algorithm that sorts objects of any kind mapped to integer keys, or else to keys in one-to-one correspondence with some subset of the integers (e.g. alphabet letters). However, it is suitable for direct use, viz. not just as a subroutine of another sorting algorithm (e.g. radix sort), only if the key range is not significantly larger than the number of the objects to be sorted. This paper describes a tail-recursive generalization of counting sort making use of a bounded number of counters, suitable for direct use in case of a large, or even infinite key range of any kind, subject to the only constraint of being a subset of an arbitrary linear order. After performing a pen-and-paper analysis of how such algorithm has to be designed to maximize its efficiency, this paper formalizes the resulting generalized counting sort (GCsort) algorithm and then formally proves its correctness properties, namely that (a) the counters' number is maximized never exceeding the fixed upper bound, (b) objects are conserved, (c) objects get sorted, and (d) the algorithm is stable. [Poincare_Bendixson] title = The Poincaré-Bendixson Theorem author = Fabian Immler , Yong Kiam Tan topic = Mathematics/Analysis date = 2019-12-18 notify = fimmler@cs.cmu.edu, yongkiat@cs.cmu.edu abstract = The Poincaré-Bendixson theorem is a classical result in the study of (continuous) dynamical systems. Colloquially, it restricts the possible behaviors of planar dynamical systems: such systems cannot be chaotic. In practice, it is a useful tool for proving the existence of (limiting) periodic behavior in planar systems. The theorem is an interesting and challenging benchmark for formalized mathematics because proofs in the literature rely on geometric sketches and only hint at symmetric cases. It also requires a substantial background of mathematical theories, e.g., the Jordan curve theorem, real analysis, ordinary differential equations, and limiting (long-term) behavior of dynamical systems. [Isabelle_C] title = Isabelle/C author = Frédéric Tuong , Burkhart Wolff topic = Computer science/Programming languages/Language definitions, Computer science/Semantics, Tools date = 2019-10-22 notify = tuong@users.gforge.inria.fr, wolff@lri.fr abstract = We present a framework for C code in C11 syntax deeply integrated into the Isabelle/PIDE development environment. Our framework provides an abstract interface for verification back-ends to be plugged-in independently. Thus, various techniques such as deductive program verification or white-box testing can be applied to the same source, which is part of an integrated PIDE document model. Semantic back-ends are free to choose the supported C fragment and its semantics. In particular, they can differ on the chosen memory model or the specification mechanism for framing conditions. Our framework supports semantic annotations of C sources in the form of comments. Annotations serve to locally control back-end settings, and can express the term focus to which an annotation refers. Both the logical and the syntactic context are available when semantic annotations are evaluated. As a consequence, a formula in an annotation can refer both to HOL or C variables. Our approach demonstrates the degree of maturity and expressive power the Isabelle/PIDE sub-system has achieved in recent years. Our integration technique employs Lex and Yacc style grammars to ensure efficient deterministic parsing. This is the core-module of Isabelle/C; the AFP package for Clean and Clean_wrapper as well as AutoCorres and AutoCorres_wrapper (available via git) are applications of this front-end. [Zeta_3_Irrational] title = The Irrationality of ζ(3) author = Manuel Eberl topic = Mathematics/Number theory date = 2019-12-27 notify = manuel.eberl@tum.de abstract =

This article provides a formalisation of Beukers's straightforward analytic proof that ζ(3) is irrational. This was first proven by Apéry (which is why this result is also often called ‘Apéry's Theorem’) using a more algebraic approach. This formalisation follows Filaseta's presentation of Beukers's proof.

[Hybrid_Logic] title = Formalizing a Seligman-Style Tableau System for Hybrid Logic author = Asta Halkjær From topic = Logic/General logic/Modal logic date = 2019-12-20 notify = ahfrom@dtu.dk abstract = This work is a formalization of soundness and completeness proofs for a Seligman-style tableau system for hybrid logic. The completeness result is obtained via a synthetic approach using maximally consistent sets of tableau blocks. The formalization differs from previous work in a few ways. First, to avoid the need to backtrack in the construction of a tableau, the formalized system has no unnamed initial segment, and therefore no Name rule. Second, I show that the full Bridge rule is admissible in the system. Third, I start from rules restricted to only extend the branch with new formulas, including only witnessing diamonds that are not already witnessed, and show that the unrestricted rules are admissible. Similarly, I start from simpler versions of the @-rules and show that these are sufficient. The GoTo rule is restricted using a notion of potential such that each application consumes potential and potential is earned through applications of the remaining rules. I show that if a branch can be closed then it can be closed starting from a single unit. Finally, Nom is restricted by a fixed set of allowed nominals. The resulting system should be terminating. extra-history = Change history: [2020-06-03]: The fully restricted system has been shown complete by updating the synthetic completeness proof. [Bicategory] title = Bicategories author = Eugene W. Stark topic = Mathematics/Category theory date = 2020-01-06 notify = stark@cs.stonybrook.edu abstract =

Taking as a starting point the author's previous work on developing aspects of category theory in Isabelle/HOL, this article gives a compatible formalization of the notion of "bicategory" and develops a framework within which formal proofs of facts about bicategories can be given. The framework includes a number of basic results, including the Coherence Theorem, the Strictness Theorem, pseudofunctors and biequivalence, and facts about internal equivalences and adjunctions in a bicategory. As a driving application and demonstration of the utility of the framework, it is used to give a formal proof of a theorem, due to Carboni, Kasangian, and Street, that characterizes up to biequivalence the bicategories of spans in a category with pullbacks. The formalization effort necessitated the filling-in of many details that were not evident from the brief presentation in the original paper, as well as identifying a few minor corrections along the way.

Revisions made subsequent to the first version of this article added additional material on pseudofunctors, pseudonatural transformations, modifications, and equivalence of bicategories; the main thrust being to give a proof that a pseudofunctor is a biequivalence if and only if it can be extended to an equivalence of bicategories.

extra-history = Change history: [2020-02-15]: Move ConcreteCategory.thy from Bicategory to Category3 and use it systematically. Make other minor improvements throughout. (revision a51840d36867)
[2020-11-04]: Added new material on equivalence of bicategories, with associated changes. (revision 472cb2268826)
[2021-07-22]: Added new material: "concrete bicategories" and "bicategory of categories". (revision 49d3aa43c180)
[Subset_Boolean_Algebras] title = A Hierarchy of Algebras for Boolean Subsets author = Walter Guttmann , Bernhard Möller topic = Mathematics/Algebra date = 2020-01-31 notify = walter.guttmann@canterbury.ac.nz abstract = We present a collection of axiom systems for the construction of Boolean subalgebras of larger overall algebras. The subalgebras are defined as the range of a complement-like operation on a semilattice. This technique has been used, for example, with the antidomain operation, dynamic negation and Stone algebras. We present a common ground for these constructions based on a new equational axiomatisation of Boolean algebras. [Goodstein_Lambda] title = Implementing the Goodstein Function in λ-Calculus author = Bertram Felgenhauer topic = Logic/Rewriting date = 2020-02-21 notify = int-e@gmx.de abstract = In this formalization, we develop an implementation of the Goodstein function G in plain λ-calculus, linked to a concise, self-contained specification. The implementation works on a Church-encoded representation of countable ordinals. The initial conversion to hereditary base 2 is not covered, but the material is sufficient to compute the particular value G(16), and easily extends to other fixed arguments. [VeriComp] title = A Generic Framework for Verified Compilers author = Martin Desharnais topic = Computer science/Programming languages/Compiling date = 2020-02-10 notify = martin.desharnais@unibw.de abstract = This is a generic framework for formalizing compiler transformations. It leverages Isabelle/HOL’s locales to abstract over concrete languages and transformations. It states common definitions for language semantics, program behaviours, forward and backward simulations, and compilers. We provide generic operations, such as simulation and compiler composition, and prove general (partial) correctness theorems, resulting in reusable proof components. [Hello_World] title = Hello World author = Cornelius Diekmann , Lars Hupel topic = Computer science/Functional programming date = 2020-03-07 notify = diekmann@net.in.tum.de abstract = In this article, we present a formalization of the well-known "Hello, World!" code, including a formal framework for reasoning about IO. Our model is inspired by the handling of IO in Haskell. We start by formalizing the 🌍 and embrace the IO monad afterwards. Then we present a sample main :: IO (), followed by its proof of correctness. [WOOT_Strong_Eventual_Consistency] title = Strong Eventual Consistency of the Collaborative Editing Framework WOOT author = Emin Karayel , Edgar Gonzàlez topic = Computer science/Algorithms/Distributed date = 2020-03-25 notify = eminkarayel@google.com, edgargip@google.com, me@eminkarayel.de abstract = Commutative Replicated Data Types (CRDTs) are a promising new class of data structures for large-scale shared mutable content in applications that only require eventual consistency. The WithOut Operational Transforms (WOOT) framework is a CRDT for collaborative text editing introduced by Oster et al. (CSCW 2006) for which the eventual consistency property was verified only for a bounded model to date. We contribute a formal proof for WOOTs strong eventual consistency. [Furstenberg_Topology] title = Furstenberg's topology and his proof of the infinitude of primes author = Manuel Eberl topic = Mathematics/Number theory date = 2020-03-22 notify = manuel.eberl@tum.de abstract =

This article gives a formal version of Furstenberg's topological proof of the infinitude of primes. He defines a topology on the integers based on arithmetic progressions (or, equivalently, residue classes). Using some fairly obvious properties of this topology, the infinitude of primes is then easily obtained.

Apart from this, this topology is also fairly ‘nice’ in general: it is second countable, metrizable, and perfect. All of these (well-known) facts are formally proven, including an explicit metric for the topology given by Zulfeqarr.

[Saturation_Framework] title = A Comprehensive Framework for Saturation Theorem Proving author = Sophie Tourret topic = Logic/General logic/Mechanization of proofs date = 2020-04-09 notify = stourret@mpi-inf.mpg.de abstract = This Isabelle/HOL formalization is the companion of the technical report “A comprehensive framework for saturation theorem proving”, itself companion of the eponym IJCAR 2020 paper, written by Uwe Waldmann, Sophie Tourret, Simon Robillard and Jasmin Blanchette. It verifies a framework for formal refutational completeness proofs of abstract provers that implement saturation calculi, such as ordered resolution or superposition, and allows to model entire prover architectures in such a way that the static refutational completeness of a calculus immediately implies the dynamic refutational completeness of a prover implementing the calculus using a variant of the given clause loop. The technical report “A comprehensive framework for saturation theorem proving” is available on the Matryoshka website. The names of the Isabelle lemmas and theorems corresponding to the results in the report are indicated in the margin of the report. [Saturation_Framework_Extensions] title = Extensions to the Comprehensive Framework for Saturation Theorem Proving author = Jasmin Blanchette , Sophie Tourret topic = Logic/General logic/Mechanization of proofs date = 2020-08-25 notify = jasmin.blanchette@gmail.com abstract = This Isabelle/HOL formalization extends the AFP entry Saturation_Framework with the following contributions:
  • an application of the framework to prove Bachmair and Ganzinger's resolution prover RP refutationally complete, which was formalized in a more ad hoc fashion by Schlichtkrull et al. in the AFP entry Ordered_Resultion_Prover;
  • generalizations of various basic concepts formalized by Schlichtkrull et al., which were needed to verify RP and could be useful to formalize other calculi, such as superposition;
  • alternative proofs of fairness (and hence saturation and ultimately refutational completeness) for the given clause procedures GC and LGC, based on invariance.
[MFODL_Monitor_Optimized] title = Formalization of an Optimized Monitoring Algorithm for Metric First-Order Dynamic Logic with Aggregations author = Thibault Dardinier<>, Lukas Heimes<>, Martin Raszyk , Joshua Schneider , Dmitriy Traytel topic = Computer science/Algorithms, Logic/General logic/Modal logic, Computer science/Automata and formal languages date = 2020-04-09 notify = martin.raszyk@inf.ethz.ch, joshua.schneider@inf.ethz.ch, traytel@inf.ethz.ch abstract = A monitor is a runtime verification tool that solves the following problem: Given a stream of time-stamped events and a policy formulated in a specification language, decide whether the policy is satisfied at every point in the stream. We verify the correctness of an executable monitor for specifications given as formulas in metric first-order dynamic logic (MFODL), which combines the features of metric first-order temporal logic (MFOTL) and metric dynamic logic. Thus, MFODL supports real-time constraints, first-order parameters, and regular expressions. Additionally, the monitor supports aggregation operations such as count and sum. This formalization, which is described in a forthcoming paper at IJCAR 2020, significantly extends previous work on a verified monitor for MFOTL. Apart from the addition of regular expressions and aggregations, we implemented multi-way joins and a specialized sliding window algorithm to further optimize the monitor. extra-history = Change history: [2021-10-19]: corrected a mistake in the calculation of median aggregations (reported by Nicolas Kaletsch, revision 02b14c9bf3da)
[Sliding_Window_Algorithm] title = Formalization of an Algorithm for Greedily Computing Associative Aggregations on Sliding Windows author = Lukas Heimes<>, Dmitriy Traytel , Joshua Schneider<> topic = Computer science/Algorithms date = 2020-04-10 notify = heimesl@student.ethz.ch, traytel@inf.ethz.ch, joshua.schneider@inf.ethz.ch abstract = Basin et al.'s sliding window algorithm (SWA) is an algorithm for combining the elements of subsequences of a sequence with an associative operator. It is greedy and minimizes the number of operator applications. We formalize the algorithm and verify its functional correctness. We extend the algorithm with additional operations and provide an alternative interface to the slide operation that does not require the entire input sequence. [Lucas_Theorem] title = Lucas's Theorem author = Chelsea Edmonds topic = Mathematics/Number theory date = 2020-04-07 notify = cle47@cam.ac.uk abstract = This work presents a formalisation of a generating function proof for Lucas's theorem. We first outline extensions to the existing Formal Power Series (FPS) library, including an equivalence relation for coefficients modulo n, an alternate binomial theorem statement, and a formalised proof of the Freshman's dream (mod p) lemma. The second part of the work presents the formal proof of Lucas's Theorem. Working backwards, the formalisation first proves a well known corollary of the theorem which is easier to formalise, and then applies induction to prove the original theorem statement. The proof of the corollary aims to provide a good example of a formalised generating function equivalence proof using the FPS library. The final theorem statement is intended to be integrated into the formalised proof of Hilbert's 10th Problem. [ADS_Functor] title = Authenticated Data Structures As Functors author = Andreas Lochbihler , Ognjen Marić topic = Computer science/Data structures date = 2020-04-16 notify = andreas.lochbihler@digitalasset.com, mail@andreas-lochbihler.de abstract = Authenticated data structures allow several systems to convince each other that they are referring to the same data structure, even if each of them knows only a part of the data structure. Using inclusion proofs, knowledgeable systems can selectively share their knowledge with other systems and the latter can verify the authenticity of what is being shared. In this article, we show how to modularly define authenticated data structures, their inclusion proofs, and operations thereon as datatypes in Isabelle/HOL, using a shallow embedding. Modularity allows us to construct complicated trees from reusable building blocks, which we call Merkle functors. Merkle functors include sums, products, and function spaces and are closed under composition and least fixpoints. As a practical application, we model the hierarchical transactions of Canton, a practical interoperability protocol for distributed ledgers, as authenticated data structures. This is a first step towards formalizing the Canton protocol and verifying its integrity and security guarantees. [Power_Sum_Polynomials] title = Power Sum Polynomials author = Manuel Eberl topic = Mathematics/Algebra date = 2020-04-24 notify = manuel@pruvisto.org abstract =

This article provides a formalisation of the symmetric multivariate polynomials known as power sum polynomials. These are of the form pn(X1,…, Xk) = X1n + … + Xkn. A formal proof of the Girard–Newton Theorem is also given. This theorem relates the power sum polynomials to the elementary symmetric polynomials sk in the form of a recurrence relation (-1)k k sk = ∑i∈[0,k) (-1)i si pk-i .

As an application, this is then used to solve a generalised form of a puzzle given as an exercise in Dummit and Foote's Abstract Algebra: For k complex unknowns x1, …, xk, define pj := x1j + … + xkj. Then for each vector a ∈ ℂk, show that there is exactly one solution to the system p1 = a1, …, pk = ak up to permutation of the xi and determine the value of pi for i>k.

[Formal_Puiseux_Series] title = Formal Puiseux Series author = Manuel Eberl topic = Mathematics/Algebra date = 2021-02-17 notify = manuel@pruvisto.org abstract =

Formal Puiseux series are generalisations of formal power series and formal Laurent series that also allow for fractional exponents. They have the following general form: \[\sum_{i=N}^\infty a_{i/d} X^{i/d}\] where N is an integer and d is a positive integer.

This entry defines these series including their basic algebraic properties. Furthermore, it proves the Newton–Puiseux Theorem, namely that the Puiseux series over an algebraically closed field of characteristic 0 are also algebraically closed.

[Gaussian_Integers] title = Gaussian Integers author = Manuel Eberl topic = Mathematics/Number theory date = 2020-04-24 notify = manuel@pruvisto.org abstract =

The Gaussian integers are the subring ℤ[i] of the complex numbers, i. e. the ring of all complex numbers with integral real and imaginary part. This article provides a definition of this ring as well as proofs of various basic properties, such as that they form a Euclidean ring and a full classification of their primes. An executable (albeit not very efficient) factorisation algorithm is also provided.

Lastly, this Gaussian integer formalisation is used in two short applications:

  1. The characterisation of all positive integers that can be written as sums of two squares
  2. Euclid's formula for primitive Pythagorean triples

While elementary proofs for both of these are already available in the AFP, the theory of Gaussian integers provides more concise proofs and a more high-level view.

[Forcing] title = Formalization of Forcing in Isabelle/ZF author = Emmanuel Gunther , Miguel Pagano , Pedro Sánchez Terraf topic = Logic/Set theory date = 2020-05-06 notify = gunther@famaf.unc.edu.ar, pagano@famaf.unc.edu.ar, sterraf@famaf.unc.edu.ar abstract = We formalize the theory of forcing in the set theory framework of Isabelle/ZF. Under the assumption of the existence of a countable transitive model of ZFC, we construct a proper generic extension and show that the latter also satisfies ZFC. [Delta_System_Lemma] title = Cofinality and the Delta System Lemma author = Pedro Sánchez Terraf topic = Mathematics/Combinatorics, Logic/Set theory date = 2020-12-27 notify = sterraf@famaf.unc.edu.ar abstract = We formalize the basic results on cofinality of linearly ordered sets and ordinals and Šanin’s Lemma for uncountable families of finite sets. This last result is used to prove the countable chain condition for Cohen posets. We work in the set theory framework of Isabelle/ZF, using the Axiom of Choice as needed. [Recursion-Addition] title = Recursion Theorem in ZF author = Georgy Dunaev topic = Logic/Set theory date = 2020-05-11 notify = georgedunaev@gmail.com abstract = This document contains a proof of the recursion theorem. This is a mechanization of the proof of the recursion theorem from the text Introduction to Set Theory, by Karel Hrbacek and Thomas Jech. This implementation may be used as the basis for a model of Peano arithmetic in ZF. While recursion and the natural numbers are already available in Isabelle/ZF, this clean development is much easier to follow. [LTL_Normal_Form] title = An Efficient Normalisation Procedure for Linear Temporal Logic: Isabelle/HOL Formalisation author = Salomon Sickert topic = Computer science/Automata and formal languages, Logic/General logic/Temporal logic date = 2020-05-08 notify = s.sickert@tum.de abstract = In the mid 80s, Lichtenstein, Pnueli, and Zuck proved a classical theorem stating that every formula of Past LTL (the extension of LTL with past operators) is equivalent to a formula of the form $\bigwedge_{i=1}^n \mathbf{G}\mathbf{F} \varphi_i \vee \mathbf{F}\mathbf{G} \psi_i$, where $\varphi_i$ and $\psi_i$ contain only past operators. Some years later, Chang, Manna, and Pnueli built on this result to derive a similar normal form for LTL. Both normalisation procedures have a non-elementary worst-case blow-up, and follow an involved path from formulas to counter-free automata to star-free regular expressions and back to formulas. We improve on both points. We present an executable formalisation of a direct and purely syntactic normalisation procedure for LTL yielding a normal form, comparable to the one by Chang, Manna, and Pnueli, that has only a single exponential blow-up. [Matrices_for_ODEs] title = Matrices for ODEs author = Jonathan Julian Huerta y Munive topic = Mathematics/Analysis, Mathematics/Algebra date = 2020-04-19 notify = jonjulian23@gmail.com abstract = Our theories formalise various matrix properties that serve to establish existence, uniqueness and characterisation of the solution to affine systems of ordinary differential equations (ODEs). In particular, we formalise the operator and maximum norm of matrices. Then we use them to prove that square matrices form a Banach space, and in this setting, we show an instance of Picard-Lindelöf’s theorem for affine systems of ODEs. Finally, we use this formalisation to verify three simple hybrid programs. [Irrational_Series_Erdos_Straus] title = Irrationality Criteria for Series by Erdős and Straus author = Angeliki Koutsoukou-Argyraki , Wenda Li topic = Mathematics/Number theory, Mathematics/Analysis date = 2020-05-12 notify = ak2110@cam.ac.uk, wl302@cam.ac.uk, liwenda1990@hotmail.com abstract = We formalise certain irrationality criteria for infinite series of the form: \[\sum_{n=1}^\infty \frac{b_n}{\prod_{i=1}^n a_i} \] where $\{b_n\}$ is a sequence of integers and $\{a_n\}$ a sequence of positive integers with $a_n >1$ for all large n. The results are due to P. Erdős and E. G. Straus [1]. In particular, we formalise Theorem 2.1, Corollary 2.10 and Theorem 3.1. The latter is an application of Theorem 2.1 involving the prime numbers. [Knuth_Bendix_Order] title = A Formalization of Knuth–Bendix Orders author = Christian Sternagel , René Thiemann topic = Logic/Rewriting date = 2020-05-13 notify = c.sternagel@gmail.com, rene.thiemann@uibk.ac.at abstract = We define a generalized version of Knuth–Bendix orders, including subterm coefficient functions. For these orders we formalize several properties such as strong normalization, the subterm property, closure properties under substitutions and contexts, as well as ground totality. [Stateful_Protocol_Composition_and_Typing] title = Stateful Protocol Composition and Typing author = Andreas V. Hess , Sebastian Mödersheim , Achim D. Brucker topic = Computer science/Security date = 2020-04-08 notify = avhe@dtu.dk, andreasvhess@gmail.com, samo@dtu.dk, brucker@spamfence.net, andschl@dtu.dk abstract = We provide in this AFP entry several relative soundness results for security protocols. In particular, we prove typing and compositionality results for stateful protocols (i.e., protocols with mutable state that may span several sessions), and that focuses on reachability properties. Such results are useful to simplify protocol verification by reducing it to a simpler problem: Typing results give conditions under which it is safe to verify a protocol in a typed model where only "well-typed" attacks can occur whereas compositionality results allow us to verify a composed protocol by only verifying the component protocols in isolation. The conditions on the protocols under which the results hold are furthermore syntactic in nature allowing for full automation. The foundation presented here is used in another entry to provide fully automated and formalized security proofs of stateful protocols. [Automated_Stateful_Protocol_Verification] title = Automated Stateful Protocol Verification author = Andreas V. Hess , Sebastian Mödersheim , Achim D. Brucker , Anders Schlichtkrull topic = Computer science/Security, Tools date = 2020-04-08 notify = avhe@dtu.dk, andreasvhess@gmail.com, samo@dtu.dk, brucker@spamfence.net, andschl@dtu.dk abstract = In protocol verification we observe a wide spectrum from fully automated methods to interactive theorem proving with proof assistants like Isabelle/HOL. In this AFP entry, we present a fully-automated approach for verifying stateful security protocols, i.e., protocols with mutable state that may span several sessions. The approach supports reachability goals like secrecy and authentication. We also include a simple user-friendly transaction-based protocol specification language that is embedded into Isabelle. [Smith_Normal_Form] title = A verified algorithm for computing the Smith normal form of a matrix author = Jose Divasón topic = Mathematics/Algebra, Computer science/Algorithms/Mathematical date = 2020-05-23 notify = jose.divason@unirioja.es abstract = This work presents a formal proof in Isabelle/HOL of an algorithm to transform a matrix into its Smith normal form, a canonical matrix form, in a general setting: the algorithm is parameterized by operations to prove its existence over elementary divisor rings, while execution is guaranteed over Euclidean domains. We also provide a formal proof on some results about the generality of this algorithm as well as the uniqueness of the Smith normal form. Since Isabelle/HOL does not feature dependent types, the development is carried out switching conveniently between two different existing libraries: the Hermite normal form (based on HOL Analysis) and the Jordan normal form AFP entries. This permits to reuse results from both developments and it is done by means of the lifting and transfer package together with the use of local type definitions. [Nash_Williams] title = The Nash-Williams Partition Theorem author = Lawrence C. Paulson topic = Mathematics/Combinatorics date = 2020-05-16 notify = lp15@cam.ac.uk abstract = In 1965, Nash-Williams discovered a generalisation of the infinite form of Ramsey's theorem. Where the latter concerns infinite sets of n-element sets for some fixed n, the Nash-Williams theorem concerns infinite sets of finite sets (or lists) subject to a “no initial segment” condition. The present formalisation follows a monograph on Ramsey Spaces by Todorčević. [Safe_Distance] title = A Formally Verified Checker of the Safe Distance Traffic Rules for Autonomous Vehicles author = Albert Rizaldi , Fabian Immler topic = Computer science/Algorithms/Mathematical, Mathematics/Physics date = 2020-06-01 notify = albert.rizaldi@ntu.edu.sg, fimmler@andrew.cmu.edu, martin.rau@tum.de abstract = The Vienna Convention on Road Traffic defines the safe distance traffic rules informally. This could make autonomous vehicle liable for safe-distance-related accidents because there is no clear definition of how large a safe distance is. We provide a formally proven prescriptive definition of a safe distance, and checkers which can decide whether an autonomous vehicle is obeying the safe distance rule. Not only does our work apply to the domain of law, but it also serves as a specification for autonomous vehicle manufacturers and for online verification of path planners. [Relational_Paths] title = Relational Characterisations of Paths author = Walter Guttmann , Peter Höfner topic = Mathematics/Graph theory date = 2020-07-13 notify = walter.guttmann@canterbury.ac.nz, peter@hoefner-online.de abstract = Binary relations are one of the standard ways to encode, characterise and reason about graphs. Relation algebras provide equational axioms for a large fragment of the calculus of binary relations. Although relations are standard tools in many areas of mathematics and computing, researchers usually fall back to point-wise reasoning when it comes to arguments about paths in a graph. We present a purely algebraic way to specify different kinds of paths in Kleene relation algebras, which are relation algebras equipped with an operation for reflexive transitive closure. We study the relationship between paths with a designated root vertex and paths without such a vertex. Since we stay in first-order logic this development helps with mechanising proofs. To demonstrate the applicability of the algebraic framework we verify the correctness of three basic graph algorithms. [Amicable_Numbers] title = Amicable Numbers author = Angeliki Koutsoukou-Argyraki topic = Mathematics/Number theory date = 2020-08-04 notify = ak2110@cam.ac.uk abstract = This is a formalisation of Amicable Numbers, involving some relevant material including Euler's sigma function, some relevant definitions, results and examples as well as rules such as Thābit ibn Qurra's Rule, Euler's Rule, te Riele's Rule and Borho's Rule with breeders. [Ordinal_Partitions] title = Ordinal Partitions author = Lawrence C. Paulson topic = Mathematics/Combinatorics, Logic/Set theory date = 2020-08-03 notify = lp15@cam.ac.uk abstract = The theory of partition relations concerns generalisations of Ramsey's theorem. For any ordinal $\alpha$, write $\alpha \to (\alpha, m)^2$ if for each function $f$ from unordered pairs of elements of $\alpha$ into $\{0,1\}$, either there is a subset $X\subseteq \alpha$ order-isomorphic to $\alpha$ such that $f\{x,y\}=0$ for all $\{x,y\}\subseteq X$, or there is an $m$ element set $Y\subseteq \alpha$ such that $f\{x,y\}=1$ for all $\{x,y\}\subseteq Y$. (In both cases, with $\{x,y\}$ we require $x\not=y$.) In particular, the infinite Ramsey theorem can be written in this notation as $\omega \to (\omega, \omega)^2$, or if we restrict $m$ to the positive integers as above, then $\omega \to (\omega, m)^2$ for all $m$. This entry formalises Larson's proof of $\omega^\omega \to (\omega^\omega, m)^2$ along with a similar proof of a result due to Specker: $\omega^2 \to (\omega^2, m)^2$. Also proved is a necessary result by Erdős and Milner: $\omega^{1+\alpha\cdot n} \to (\omega^{1+\alpha}, 2^n)^2$. [Relational_Disjoint_Set_Forests] title = Relational Disjoint-Set Forests author = Walter Guttmann topic = Computer science/Data structures date = 2020-08-26 notify = walter.guttmann@canterbury.ac.nz abstract = We give a simple relation-algebraic semantics of read and write operations on associative arrays. The array operations seamlessly integrate with assignments in the Hoare-logic library. Using relation algebras and Kleene algebras we verify the correctness of an array-based implementation of disjoint-set forests with a naive union operation and a find operation with path compression. extra-history = Change history: [2021-06-19]: added path halving, path splitting, relational Peano structures, union by rank (revision 98c7aa03457d) [PAC_Checker] title = Practical Algebraic Calculus Checker author = Mathias Fleury , Daniela Kaufmann topic = Computer science/Algorithms date = 2020-08-31 notify = mathias.fleury@jku.at abstract = Generating and checking proof certificates is important to increase the trust in automated reasoning tools. In recent years formal verification using computer algebra became more important and is heavily used in automated circuit verification. An existing proof format which covers algebraic reasoning and allows efficient proof checking is the practical algebraic calculus (PAC). In this development, we present the verified checker Pastèque that is obtained by synthesis via the Refinement Framework. This is the formalization going with our FMCAD'20 tool presentation. [BirdKMP] title = Putting the `K' into Bird's derivation of Knuth-Morris-Pratt string matching author = Peter Gammie topic = Computer science/Functional programming date = 2020-08-25 notify = peteg42@gmail.com abstract = Richard Bird and collaborators have proposed a derivation of an intricate cyclic program that implements the Morris-Pratt string matching algorithm. Here we provide a proof of total correctness for Bird's derivation and complete it by adding Knuth's optimisation. [Extended_Finite_State_Machines] title = A Formal Model of Extended Finite State Machines author = Michael Foster , Achim D. Brucker , Ramsay G. Taylor , John Derrick topic = Computer science/Automata and formal languages date = 2020-09-07 notify = jmafoster1@sheffield.ac.uk, adbrucker@0x5f.org abstract = In this AFP entry, we provide a formalisation of extended finite state machines (EFSMs) where models are represented as finite sets of transitions between states. EFSMs execute traces to produce observable outputs. We also define various simulation and equality metrics for EFSMs in terms of traces and prove their strengths in relation to each other. Another key contribution is a framework of function definitions such that LTL properties can be phrased over EFSMs. Finally, we provide a simple example case study in the form of a drinks machine. [Extended_Finite_State_Machine_Inference] title = Inference of Extended Finite State Machines author = Michael Foster , Achim D. Brucker , Ramsay G. Taylor , John Derrick topic = Computer science/Automata and formal languages date = 2020-09-07 notify = jmafoster1@sheffield.ac.uk, adbrucker@0x5f.org abstract = In this AFP entry, we provide a formal implementation of a state-merging technique to infer extended finite state machines (EFSMs), complete with output and update functions, from black-box traces. In particular, we define the subsumption in context relation as a means of determining whether one transition is able to account for the behaviour of another. Building on this, we define the direct subsumption relation, which lifts the subsumption in context relation to EFSM level such that we can use it to determine whether it is safe to merge a given pair of transitions. Key proofs include the conditions necessary for subsumption to occur and that subsumption and direct subsumption are preorder relations. We also provide a number of different heuristics which can be used to abstract away concrete values into registers so that more states and transitions can be merged and provide proofs of the various conditions which must hold for these abstractions to subsume their ungeneralised counterparts. A Code Generator setup to create executable Scala code is also defined. [Physical_Quantities] title = A Sound Type System for Physical Quantities, Units, and Measurements author = Simon Foster , Burkhart Wolff topic = Mathematics/Physics, Computer science/Programming languages/Type systems date = 2020-10-20 notify = simon.foster@york.ac.uk, wolff@lri.fr abstract = The present Isabelle theory builds a formal model for both the International System of Quantities (ISQ) and the International System of Units (SI), which are both fundamental for physics and engineering. Both the ISQ and the SI are deeply integrated into Isabelle's type system. Quantities are parameterised by dimension types, which correspond to base vectors, and thus only quantities of the same dimension can be equated. Since the underlying "algebra of quantities" induces congruences on quantity and SI types, specific tactic support is developed to capture these. Our construction is validated by a test-set of known equivalences between both quantities and SI units. Moreover, the presented theory can be used for type-safe conversions between the SI system and others, like the British Imperial System (BIS). [Shadow_DOM] title = A Formal Model of the Document Object Model with Shadow Roots author = Achim D. Brucker , Michael Herzberg topic = Computer science/Data structures date = 2020-09-28 notify = adbrucker@0x5f.org, mail@michael-herzberg.de abstract = In this AFP entry, we extend our formalization of the core DOM with Shadow Roots. Shadow roots are a recent proposal of the web community to support a component-based development approach for client-side web applications. Shadow roots are a significant extension to the DOM standard and, as web standards are condemned to be backward compatible, such extensions often result in complex specification that may contain unwanted subtleties that can be detected by a formalization. Our Isabelle/HOL formalization is, in the sense of object-orientation, an extension of our formalization of the core DOM and enjoys the same basic properties, i.e., it is extensible, i.e., can be extended without the need of re-proving already proven properties and executable, i.e., we can generate executable code from our specification. We exploit the executability to show that our formalization complies to the official standard of the W3C, respectively, the WHATWG. [DOM_Components] title = A Formalization of Web Components author = Achim D. Brucker , Michael Herzberg topic = Computer science/Data structures date = 2020-09-28 notify = adbrucker@0x5f.org, mail@michael-herzberg.de abstract = While the DOM with shadow trees provide the technical basis for defining web components, the DOM standard neither defines the concept of web components nor specifies the safety properties that web components should guarantee. Consequently, the standard also does not discuss how or even if the methods for modifying the DOM respect component boundaries. In AFP entry, we present a formally verified model of web components and define safety properties which ensure that different web components can only interact with each other using well-defined interfaces. Moreover, our verification of the application programming interface (API) of the DOM revealed numerous invariants that implementations of the DOM API need to preserve to ensure the integrity of components. [Interpreter_Optimizations] title = Inline Caching and Unboxing Optimization for Interpreters author = Martin Desharnais topic = Computer science/Programming languages/Misc date = 2020-12-07 notify = martin.desharnais@unibw.de abstract = This Isabelle/HOL formalization builds on the VeriComp entry of the Archive of Formal Proofs to provide the following contributions:
  • an operational semantics for a realistic virtual machine (Std) for dynamically typed programming languages;
  • the formalization of an inline caching optimization (Inca), a proof of bisimulation with (Std), and a compilation function;
  • the formalization of an unboxing optimization (Ubx), a proof of bisimulation with (Inca), and a simple compilation function.
This formalization was described in the CPP 2021 paper Towards Efficient and Verified Virtual Machines for Dynamic Languages extra-history = Change history: [2021-06-14]: refactored function definitions to contain explicit basic blocks
[2021-06-25]: proved conditional completeness of compilation
[Isabelle_Marries_Dirac] title = Isabelle Marries Dirac: a Library for Quantum Computation and Quantum Information author = Anthony Bordg , Hanna Lachnitt, Yijun He topic = Computer science/Algorithms/Quantum computing, Mathematics/Physics/Quantum information date = 2020-11-22 notify = apdb3@cam.ac.uk, lachnitt@stanford.edu abstract = This work is an effort to formalise some quantum algorithms and results in quantum information theory. Formal methods being critical for the safety and security of algorithms and protocols, we foresee their widespread use for quantum computing in the future. We have developed a large library for quantum computing in Isabelle based on a matrix representation for quantum circuits, successfully formalising the no-cloning theorem, quantum teleportation, Deutsch's algorithm, the Deutsch-Jozsa algorithm and the quantum Prisoner's Dilemma. [Projective_Measurements] title = Quantum projective measurements and the CHSH inequality author = Mnacho Echenim topic = Computer science/Algorithms/Quantum computing, Mathematics/Physics/Quantum information date = 2021-03-03 notify = mnacho.echenim@univ-grenoble-alpes.fr abstract = This work contains a formalization of quantum projective measurements, also known as von Neumann measurements, which are based on elements of spectral theory. We also formalized the CHSH inequality, an inequality involving expectations in a probability space that is violated by quantum measurements, thus proving that quantum mechanics cannot be modeled with an underlying local hidden-variable theory. [Finite-Map-Extras] title = Finite Map Extras author = Javier Díaz topic = Computer science/Data structures date = 2020-10-12 notify = javier.diaz.manzi@gmail.com abstract = This entry includes useful syntactic sugar, new operators and functions, and their associated lemmas for finite maps which currently are not present in the standard Finite_Map theory. [Relational_Minimum_Spanning_Trees] title = Relational Minimum Spanning Tree Algorithms author = Walter Guttmann , Nicolas Robinson-O'Brien<> topic = Computer science/Algorithms/Graph date = 2020-12-08 notify = walter.guttmann@canterbury.ac.nz abstract = We verify the correctness of Prim's, Kruskal's and Borůvka's minimum spanning tree algorithms based on algebras for aggregation and minimisation. [Topological_Semantics] title = Topological semantics for paraconsistent and paracomplete logics author = David Fuenmayor topic = Logic/General logic date = 2020-12-17 notify = davfuenmayor@gmail.com abstract = We introduce a generalized topological semantics for paraconsistent and paracomplete logics by drawing upon early works on topological Boolean algebras (cf. works by Kuratowski, Zarycki, McKinsey & Tarski, etc.). In particular, this work exemplarily illustrates the shallow semantical embeddings approach (SSE) employing the proof assistant Isabelle/HOL. By means of the SSE technique we can effectively harness theorem provers, model finders and 'hammers' for reasoning with quantified non-classical logics. [CSP_RefTK] title = The HOL-CSP Refinement Toolkit author = Safouan Taha , Burkhart Wolff , Lina Ye topic = Computer science/Concurrency/Process calculi, Computer science/Semantics date = 2020-11-19 notify = wolff@lri.fr abstract = We use a formal development for CSP, called HOL-CSP2.0, to analyse a family of refinement notions, comprising classic and new ones. This analysis enables to derive a number of properties that allow to deepen the understanding of these notions, in particular with respect to specification decomposition principles for the case of infinite sets of events. The established relations between the refinement relations help to clarify some obscure points in the CSP literature, but also provide a weapon for shorter refinement proofs. Furthermore, we provide a framework for state-normalisation allowing to formally reason on parameterised process architectures. As a result, we have a modern environment for formal proofs of concurrent systems that allow for the combination of general infinite processes with locally finite ones in a logically safe way. We demonstrate these verification-techniques for classical, generalised examples: The CopyBuffer for arbitrary data and the Dijkstra's Dining Philosopher Problem of arbitrary size. [Hood_Melville_Queue] title = Hood-Melville Queue author = Alejandro Gómez-Londoño topic = Computer science/Data structures date = 2021-01-18 notify = nipkow@in.tum.de abstract = This is a verified implementation of a constant time queue. The original design is due to Hood and Melville. This formalization follows the presentation in Purely Functional Data Structuresby Okasaki. [JinjaDCI] title = JinjaDCI: a Java semantics with dynamic class initialization author = Susannah Mansky topic = Computer science/Programming languages/Language definitions date = 2021-01-11 notify = sjohnsn2@illinois.edu, susannahej@gmail.com abstract = We extend Jinja to include static fields, methods, and instructions, and dynamic class initialization, based on the Java SE 8 specification. This includes extension of definitions and proofs. This work is partially described in Mansky and Gunter's paper at CPP 2019 and Mansky's doctoral thesis (UIUC, 2020). [Blue_Eyes] title = Solution to the xkcd Blue Eyes puzzle author = Jakub Kądziołka topic = Logic/General logic/Logics of knowledge and belief date = 2021-01-30 notify = kuba@kadziolka.net abstract = In a puzzle published by Randall Munroe, perfect logicians forbidden from communicating are stranded on an island, and may only leave once they have figured out their own eye color. We present a method of modeling the behavior of perfect logicians and formalize a solution of the puzzle. [Laws_of_Large_Numbers] title = The Laws of Large Numbers author = Manuel Eberl topic = Mathematics/Probability theory date = 2021-02-10 notify = manuel@pruvisto.org abstract =

The Law of Large Numbers states that, informally, if one performs a random experiment $X$ many times and takes the average of the results, that average will be very close to the expected value $E[X]$.

More formally, let $(X_i)_{i\in\mathbb{N}}$ be a sequence of independently identically distributed random variables whose expected value $E[X_1]$ exists. Denote the running average of $X_1, \ldots, X_n$ as $\overline{X}_n$. Then:

  • The Weak Law of Large Numbers states that $\overline{X}_{n} \longrightarrow E[X_1]$ in probability for $n\to\infty$, i.e. $\mathcal{P}(|\overline{X}_{n} - E[X_1]| > \varepsilon) \longrightarrow 0$ as $n\to\infty$ for any $\varepsilon > 0$.
  • The Strong Law of Large Numbers states that $\overline{X}_{n} \longrightarrow E[X_1]$ almost surely for $n\to\infty$, i.e. $\mathcal{P}(\overline{X}_{n} \longrightarrow E[X_1]) = 1$.

In this entry, I formally prove the strong law and from it the weak law. The approach used for the proof of the strong law is a particularly quick and slick one based on ergodic theory, which was formalised by Gouëzel in another AFP entry.

[BTree] title = A Verified Imperative Implementation of B-Trees author = Niels Mündler topic = Computer science/Data structures date = 2021-02-24 notify = n.muendler@tum.de abstract = In this work, we use the interactive theorem prover Isabelle/HOL to verify an imperative implementation of the classical B-tree data structure invented by Bayer and McCreight [ACM 1970]. The implementation supports set membership, insertion and deletion queries with efficient binary search for intra-node navigation. This is accomplished by first specifying the structure abstractly in the functional modeling language HOL and proving functional correctness. Using manual refinement, we derive an imperative implementation in Imperative/HOL. We show the validity of this refinement using the separation logic utilities from the Isabelle Refinement Framework . The code can be exported to the programming languages SML, OCaml and Scala. We examine the runtime of all operations indirectly by reproducing results of the logarithmic relationship between height and the number of nodes. The results are discussed in greater detail in the corresponding Bachelor's Thesis. extra-history = Change history: [2021-05-02]: Add implementation and proof of correctness of imperative deletion operations. Further add the option to export code to OCaml.
[Sunflowers] title = The Sunflower Lemma of Erdős and Rado author = René Thiemann topic = Mathematics/Combinatorics date = 2021-02-25 notify = rene.thiemann@uibk.ac.at abstract = We formally define sunflowers and provide a formalization of the sunflower lemma of Erdős and Rado: whenever a set of size-k-sets has a larger cardinality than (r - 1)k · k!, then it contains a sunflower of cardinality r. [Mereology] title = Mereology author = Ben Blumson topic = Logic/Philosophical aspects date = 2021-03-01 notify = benblumson@gmail.com abstract = We use Isabelle/HOL to verify elementary theorems and alternative axiomatizations of classical extensional mereology. [Modular_arithmetic_LLL_and_HNF_algorithms] title = Two algorithms based on modular arithmetic: lattice basis reduction and Hermite normal form computation author = Ralph Bottesch <>, Jose Divasón , René Thiemann topic = Computer science/Algorithms/Mathematical date = 2021-03-12 notify = rene.thiemann@uibk.ac.at abstract = We verify two algorithms for which modular arithmetic plays an essential role: Storjohann's variant of the LLL lattice basis reduction algorithm and Kopparty's algorithm for computing the Hermite normal form of a matrix. To do this, we also formalize some facts about the modulo operation with symmetric range. Our implementations are based on the original papers, but are otherwise efficient. For basis reduction we formalize two versions: one that includes all of the optimizations/heuristics from Storjohann's paper, and one excluding a heuristic that we observed to often decrease efficiency. We also provide a fast, self-contained certifier for basis reduction, based on the efficient Hermite normal form algorithm. [Constructive_Cryptography_CM] title = Constructive Cryptography in HOL: the Communication Modeling Aspect author = Andreas Lochbihler , S. Reza Sefidgar <> topic = Computer science/Security/Cryptography, Mathematics/Probability theory date = 2021-03-17 notify = mail@andreas-lochbihler.de, reza.sefidgar@inf.ethz.ch abstract = Constructive Cryptography (CC) [ICS 2011, TOSCA 2011, TCC 2016] introduces an abstract approach to composable security statements that allows one to focus on a particular aspect of security proofs at a time. Instead of proving the properties of concrete systems, CC studies system classes, i.e., the shared behavior of similar systems, and their transformations. Modeling of systems communication plays a crucial role in composability and reusability of security statements; yet, this aspect has not been studied in any of the existing CC results. We extend our previous CC formalization [Constructive_Cryptography, CSF 2019] with a new semantic domain called Fused Resource Templates (FRT) that abstracts over the systems communication patterns in CC proofs. This widens the scope of cryptography proof formalizations in the CryptHOL library [CryptHOL, ESOP 2016, J Cryptol 2020]. This formalization is described in Abstract Modeling of Systems Communication in Constructive Cryptography using CryptHOL. [IFC_Tracking] title = Information Flow Control via Dependency Tracking author = Benedikt Nordhoff topic = Computer science/Security date = 2021-04-01 notify = b.n@wwu.de abstract = We provide a characterisation of how information is propagated by program executions based on the tracking data and control dependencies within executions themselves. The characterisation might be used for deriving approximative safety properties to be targeted by static analyses or checked at runtime. We utilise a simple yet versatile control flow graph model as a program representation. As our model is not assumed to be finite it can be instantiated for a broad class of programs. The targeted security property is indistinguishable security where executions produce sequences of observations and only non-terminating executions are allowed to drop a tail of those. A very crude approximation of our characterisation is slicing based on program dependence graphs, which we use as a minimal example and derive a corresponding soundness result. For further details and applications refer to the authors upcoming dissertation. [Grothendieck_Schemes] title = Grothendieck's Schemes in Algebraic Geometry author = Anthony Bordg , Lawrence Paulson , Wenda Li topic = Mathematics/Algebra, Mathematics/Geometry date = 2021-03-29 notify = apdb3@cam.ac.uk, lp15@cam.ac.uk abstract = We formalize mainstream structures in algebraic geometry culminating in Grothendieck's schemes: presheaves of rings, sheaves of rings, ringed spaces, locally ringed spaces, affine schemes and schemes. We prove that the spectrum of a ring is a locally ringed space, hence an affine scheme. Finally, we prove that any affine scheme is a scheme. [Progress_Tracking] title = Formalization of Timely Dataflow's Progress Tracking Protocol author = Matthias Brun<>, Sára Decova<>, Andrea Lattuada, Dmitriy Traytel topic = Computer science/Algorithms/Distributed date = 2021-04-13 notify = matthias.brun@inf.ethz.ch, traytel@di.ku.dk abstract = Large-scale stream processing systems often follow the dataflow paradigm, which enforces a program structure that exposes a high degree of parallelism. The Timely Dataflow distributed system supports expressive cyclic dataflows for which it offers low-latency data- and pipeline-parallel stream processing. To achieve high expressiveness and performance, Timely Dataflow uses an intricate distributed protocol for tracking the computation’s progress. We formalize this progress tracking protocol and verify its safety. Our formalization is described in detail in our forthcoming ITP'21 paper. [GaleStewart_Games] title = Gale-Stewart Games author = Sebastiaan Joosten topic = Mathematics/Games and economics date = 2021-04-23 notify = sjcjoosten@gmail.com abstract = This is a formalisation of the main result of Gale and Stewart from 1953, showing that closed finite games are determined. This property is now known as the Gale Stewart Theorem. While the original paper shows some additional theorems as well, we only formalize this main result, but do so in a somewhat general way. We formalize games of a fixed arbitrary length, including infinite length, using co-inductive lists, and show that defensive strategies exist unless the other player is winning. For closed games, defensive strategies are winning for the closed player, proving that such games are determined. For finite games, which are a special case in our formalisation, all games are closed. [Metalogic_ProofChecker] title = Isabelle's Metalogic: Formalization and Proof Checker author = Tobias Nipkow , Simon Roßkopf topic = Logic/General logic date = 2021-04-27 notify = rosskops@in.tum.de abstract = In this entry we formalize Isabelle's metalogic in Isabelle/HOL. Furthermore, we define a language of proof terms and an executable proof checker and prove its soundness wrt. the metalogic. The formalization is intentionally kept close to the Isabelle implementation(for example using de Brujin indices) to enable easy integration of generated code with the Isabelle system without a complicated translation layer. The formalization is described in our CADE 28 paper. [Regression_Test_Selection] title = Regression Test Selection author = Susannah Mansky topic = Computer science/Algorithms date = 2021-04-30 notify = sjohnsn2@illinois.edu, susannahej@gmail.com abstract = This development provides a general definition for safe Regression Test Selection (RTS) algorithms. RTS algorithms select which tests to rerun on revised code, reducing the time required to check for newly introduced errors. An RTS algorithm is considered safe if and only if all deselected tests would have unchanged results. This definition is instantiated with two class-collection-based RTS algorithms run over the JVM as modeled by JinjaDCI. This is achieved with a general definition for Collection Semantics, small-step semantics instrumented to collect information during execution. As the RTS definition mandates safety, these instantiations include proofs of safety. This work is described in Mansky and Gunter's LSFA 2020 paper and Mansky's doctoral thesis (UIUC, 2020). [Padic_Ints] title = Hensel's Lemma for the p-adic Integers author = Aaron Crighton topic = Mathematics/Number theory date = 2021-03-23 notify = crightoa@mcmaster.ca abstract = We formalize the ring of p-adic integers within the framework of the HOL-Algebra library. The carrier of the ring is formalized as the inverse limit of quotients of the integers by powers of a fixed prime p. We define an integer-valued valuation, as well as an extended-integer valued valuation which sends 0 to the infinite element. Basic topological facts about the p-adic integers are formalized, including completeness and sequential compactness. Taylor expansions of polynomials over a commutative ring are defined, culminating in the formalization of Hensel's Lemma based on a proof due to Keith Conrad. [Combinatorics_Words] title = Combinatorics on Words Basics author = Štěpán Holub , Martin Raška<>, Štěpán Starosta topic = Computer science/Automata and formal languages date = 2021-05-24 notify = holub@karlin.mff.cuni.cz, stepan.starosta@fit.cvut.cz abstract = We formalize basics of Combinatorics on Words. This is an extension of existing theories on lists. We provide additional properties related to prefix, suffix, factor, length and rotation. The topics include prefix and suffix comparability, mismatch, word power, total and reversed morphisms, border, periods, primitivity and roots. We also formalize basic, mostly folklore results related to word equations: equidivisibility, commutation and conjugation. Slightly advanced properties include the Periodicity lemma (often cited as the Fine and Wilf theorem) and the variant of the Lyndon-Schützenberger theorem for words. We support the algebraic point of view which sees words as generators of submonoids of a free monoid. This leads to the concepts of the (free) hull, the (free) basis (or code). [Combinatorics_Words_Lyndon] title = Lyndon words author = Štěpán Holub , Štěpán Starosta topic = Computer science/Automata and formal languages date = 2021-05-24 notify = holub@karlin.mff.cuni.cz, stepan.starosta@fit.cvut.cz abstract = Lyndon words are words lexicographically minimal in their conjugacy class. We formalize their basic properties and characterizations, in particular the concepts of the longest Lyndon suffix and the Lyndon factorization. Most of the work assumes a fixed lexicographical order. Nevertheless we also define the smallest relation guaranteeing lexicographical minimality of a given word (in its conjugacy class). [Combinatorics_Words_Graph_Lemma] title = Graph Lemma author = Štěpán Holub , Štěpán Starosta topic = Computer science/Automata and formal languages date = 2021-05-24 notify = holub@karlin.mff.cuni.cz, stepan.starosta@fit.cvut.cz abstract = Graph lemma quantifies the defect effect of a system of word equations. That is, it provides an upper bound on the rank of the system. We formalize the proof based on the decomposition of a solution into its free basis. A direct application is an alternative proof of the fact that two noncommuting words form a code. [Lifting_the_Exponent] title = Lifting the Exponent author = Jakub Kądziołka topic = Mathematics/Number theory date = 2021-04-27 notify = kuba@kadziolka.net abstract = We formalize the Lifting the Exponent Lemma, which shows how to find the largest power of $p$ dividing $a^n \pm b^n$, for a prime $p$ and positive integers $a$ and $b$. The proof follows Amir Hossein Parvardi's. [IMP_Compiler] title = A Shorter Compiler Correctness Proof for Language IMP author = Pasquale Noce topic = Computer science/Programming languages/Compiling date = 2021-06-04 notify = pasquale.noce.lavoro@gmail.com abstract = This paper presents a compiler correctness proof for the didactic imperative programming language IMP, introduced in Nipkow and Klein's book on formal programming language semantics (version of March 2021), whose size is just two thirds of the book's proof in the number of formal text lines. As such, it promises to constitute a further enhanced reference for the formal verification of compilers meant for larger, real-world programming languages. The presented proof does not depend on language determinism, so that the proposed approach can be applied to non-deterministic languages as well. As a confirmation, this paper extends IMP with an additional non-deterministic choice command, and proves compiler correctness, viz. the simulation of compiled code execution by source code, for such extended language. [Public_Announcement_Logic] title = Public Announcement Logic author = Asta Halkjær From topic = Logic/General logic/Logics of knowledge and belief date = 2021-06-17 notify = ahfrom@dtu.dk abstract = This work is a formalization of public announcement logic with countably many agents. It includes proofs of soundness and completeness for a variant of the axiom system PA + DIST! + NEC!. The completeness proof builds on the Epistemic Logic theory. [MiniSail] title = MiniSail - A kernel language for the ISA specification language SAIL author = Mark Wassell topic = Computer science/Programming languages/Type systems date = 2021-06-18 notify = mpwassell@gmail.com abstract = MiniSail is a kernel language for Sail, an instruction set architecture (ISA) specification language. Sail is an imperative language with a light-weight dependent type system similar to refinement type systems. From an ISA specification, the Sail compiler can generate theorem prover code and C (or OCaml) to give an executable emulator for an architecture. The idea behind MiniSail is to capture the key and novel features of Sail in terms of their syntax, typing rules and operational semantics, and to confirm that they work together by proving progress and preservation lemmas. We use the Nominal2 library to handle binding. [SpecCheck] title = SpecCheck - Specification-Based Testing for Isabelle/ML author = Kevin Kappelmann , Lukas Bulwahn , Sebastian Willenbrink topic = Tools date = 2021-07-01 notify = kevin.kappelmann@tum.de abstract = SpecCheck is a QuickCheck-like testing framework for Isabelle/ML. You can use it to write specifications for ML functions. SpecCheck then checks whether your specification holds by testing your function against a given number of generated inputs. It helps you to identify bugs by printing counterexamples on failure and provides you timing information. SpecCheck is customisable and allows you to specify your own input generators, test output formats, as well as pretty printers and shrinking functions for counterexamples among other things. [Relational_Forests] title = Relational Forests author = Walter Guttmann topic = Mathematics/Graph theory date = 2021-08-03 notify = walter.guttmann@canterbury.ac.nz abstract = We study second-order formalisations of graph properties expressed as first-order formulas in relation algebras extended with a Kleene star. The formulas quantify over relations while still avoiding quantification over elements of the base set. We formalise the property of undirected graphs being acyclic this way. This involves a study of various kinds of orientation of graphs. We also verify basic algorithms to constructively prove several second-order properties. [Fresh_Identifiers] title = Fresh identifiers author = Andrei Popescu , Thomas Bauereiss topic = Computer science/Data structures date = 2021-08-16 notify = thomas@bauereiss.name, a.popescu@sheffield.ac.uk abstract = This entry defines a type class with an operator returning a fresh identifier, given a set of already used identifiers and a preferred identifier. The entry provides a default instantiation for any infinite type, as well as executable instantiations for natural numbers and strings. [CoCon] title = CoCon: A Confidentiality-Verified Conference Management System author = Andrei Popescu , Peter Lammich , Thomas Bauereiss topic = Computer science/Security date = 2021-08-16 notify = thomas@bauereiss.name, a.popescu@sheffield.ac.uk abstract = This entry contains the confidentiality verification of the (functional kernel of) the CoCon conference management system [1, 2]. The confidentiality properties refer to the documents managed by the system, namely papers, reviews, discussion logs and acceptance/rejection decisions, and also to the assignment of reviewers to papers. They have all been formulated as instances of BD Security [3, 4] and verified using the BD Security unwinding technique. [BD_Security_Compositional] title = Compositional BD Security author = Thomas Bauereiss , Andrei Popescu topic = Computer science/Security date = 2021-08-16 notify = thomas@bauereiss.name, a.popescu@sheffield.ac.uk abstract = Building on a previous AFP entry that formalizes the Bounded-Deducibility Security (BD Security) framework [1], we formalize compositionality and transport theorems for information flow security. These results allow lifting BD Security properties from individual components specified as transition systems, to a composition of systems specified as communicating products of transition systems. The underlying ideas of these results are presented in the papers [1] and [2]. The latter paper also describes a major case study where these results have been used: on verifying the CoSMeDis distributed social media platform (itself formalized as an AFP entry that builds on this entry). [CoSMed] title = CoSMed: A confidentiality-verified social media platform author = Thomas Bauereiss , Andrei Popescu topic = Computer science/Security date = 2021-08-16 notify = thomas@bauereiss.name, a.popescu@sheffield.ac.uk abstract = This entry contains the confidentiality verification of the (functional kernel of) the CoSMed social media platform. The confidentiality properties are formalized as instances of BD Security [1, 2]. An innovation in the deployment of BD Security compared to previous work is the use of dynamic declassification triggers, incorporated as part of inductive bounds, for providing stronger guarantees that account for the repeated opening and closing of access windows. To further strengthen the confidentiality guarantees, we also prove "traceback" properties about the accessibility decisions affecting the information managed by the system. [CoSMeDis] title = CoSMeDis: A confidentiality-verified distributed social media platform author = Thomas Bauereiss , Andrei Popescu topic = Computer science/Security date = 2021-08-16 notify = thomas@bauereiss.name, a.popescu@sheffield.ac.uk abstract = This entry contains the confidentiality verification of the (functional kernel of) the CoSMeDis distributed social media platform presented in [1]. CoSMeDis is a multi-node extension the CoSMed prototype social media platform [2, 3, 4]. The confidentiality properties are formalized as instances of BD Security [5, 6]. The lifting of confidentiality properties from single nodes to the entire CoSMeDis network is performed using compositionality and transport theorems for BD Security, which are described in [1] and formalized in a separate AFP entry. [Three_Circles] title = The Theorem of Three Circles author = Fox Thomson , Wenda Li topic = Mathematics/Analysis date = 2021-08-21 notify = foxthomson0@gmail.com, wl302@cam.ac.uk abstract = The Descartes test based on Bernstein coefficients and Descartes’ rule of signs effectively (over-)approximates the number of real roots of a univariate polynomial over an interval. In this entry we formalise the theorem of three circles, which gives sufficient conditions for when the Descartes test returns 0 or 1. This is the first step for efficient root isolation. [Design_Theory] title = Combinatorial Design Theory author = Chelsea Edmonds , Lawrence Paulson topic = Mathematics/Combinatorics date = 2021-08-13 notify = cle47@cam.ac.uk abstract = Combinatorial design theory studies incidence set systems with certain balance and symmetry properties. It is closely related to hypergraph theory. This formalisation presents a general library for formal reasoning on incidence set systems, designs and their applications, including formal definitions and proofs for many key properties, operations, and theorems on the construction and existence of designs. Notably, this includes formalising t-designs, balanced incomplete block designs (BIBD), group divisible designs (GDD), pairwise balanced designs (PBD), design isomorphisms, and the relationship between graphs and designs. A locale-centric approach has been used to manage the relationships between the many different types of designs. Theorems of particular interest include the necessary conditions for existence of a BIBD, Wilson's construction on GDDs, and Bose's inequality on resolvable designs. Parts of this formalisation are explored in the paper "A Modular First Formalisation of Combinatorial Design Theory", presented at CICM 2021. [Logging_Independent_Anonymity] title = Logging-independent Message Anonymity in the Relational Method author = Pasquale Noce topic = Computer science/Security date = 2021-08-26 notify = pasquale.noce.lavoro@gmail.com abstract = In the context of formal cryptographic protocol verification, logging-independent message anonymity is the property for a given message to remain anonymous despite the attacker's capability of mapping messages of that sort to agents based on some intrinsic feature of such messages, rather than by logging the messages exchanged by legitimate agents as with logging-dependent message anonymity. This paper illustrates how logging-independent message anonymity can be formalized according to the relational method for formal protocol verification by considering a real-world protocol, namely the Restricted Identification one by the BSI. This sample model is used to verify that the pseudonymous identifiers output by user identification tokens remain anonymous under the expected conditions. [Dominance_CHK] title = A data flow analysis algorithm for computing dominators author = Nan Jiang<> topic = Computer science/Programming languages/Static analysis date = 2021-09-05 notify = nanjiang@whu.edu.cn abstract = This entry formalises the fast iterative algorithm for computing dominators due to Cooper, Harvey and Kennedy. It gives a specification of computing dominators on a control flow graph where each node refers to its reverse post order number. A semilattice of reversed-ordered list which represents dominators is built and a Kildall-style algorithm on the semilattice is defined for computing dominators. Finally the soundness and completeness of the algorithm are proved w.r.t. the specification. [Conditional_Simplification] title = Conditional Simplification author = Mihails Milehins topic = Tools date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = The article provides a collection of experimental general-purpose proof methods for the object logic Isabelle/HOL of the formal proof assistant Isabelle. The methods in the collection offer functionality that is similar to certain aspects of the functionality provided by the standard proof methods of Isabelle that combine classical reasoning and rewriting, such as the method auto, but use a different approach for rewriting. More specifically, these methods allow for the side conditions of the rewrite rules to be solved via intro-resolution. [Intro_Dest_Elim] title = IDE: Introduction, Destruction, Elimination author = Mihails Milehins topic = Tools date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = The article provides the command mk_ide for the object logic Isabelle/HOL of the formal proof assistant Isabelle. The command mk_ide enables the automated synthesis of the introduction, destruction and elimination rules from arbitrary definitions of constant predicates stated in Isabelle/HOL. [CZH_Foundations] title = Category Theory for ZFC in HOL I: Foundations: Design Patterns, Set Theory, Digraphs, Semicategories author = Mihails Milehins topic = Mathematics/Category theory, Logic/Set theory date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = This article provides a foundational framework for the formalization of category theory in the object logic ZFC in HOL of the formal proof assistant Isabelle. More specifically, this article provides a formalization of canonical set-theoretic constructions internalized in the type V associated with the ZFC in HOL, establishes a design pattern for the formalization of mathematical structures using sequences and locales, and showcases the developed infrastructure by providing formalizations of the elementary theories of digraphs and semicategories. The methodology chosen for the formalization of the theories of digraphs and semicategories (and categories in future articles) rests on the ideas that were originally expressed in the article Set-Theoretical Foundations of Category Theory written by Solomon Feferman and Georg Kreisel. Thus, in the context of this work, each of the aforementioned mathematical structures is represented as a term of the type V embedded into a stage of the von Neumann hierarchy. [CZH_Elementary_Categories] title = Category Theory for ZFC in HOL II: Elementary Theory of 1-Categories author = Mihails Milehins topic = Mathematics/Category theory date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = This article provides a formalization of the foundations of the theory of 1-categories in the object logic ZFC in HOL of the formal proof assistant Isabelle. The article builds upon the foundations that were established in the AFP entry Category Theory for ZFC in HOL I: Foundations: Design Patterns, Set Theory, Digraphs, Semicategories. [CZH_Universal_Constructions] title = Category Theory for ZFC in HOL III: Universal Constructions author = Mihails Milehins topic = Mathematics/Category theory date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = The article provides a formalization of elements of the theory of universal constructions for 1-categories (such as limits, adjoints and Kan extensions) in the object logic ZFC in HOL of the formal proof assistant Isabelle. The article builds upon the foundations established in the AFP entry Category Theory for ZFC in HOL II: Elementary Theory of 1-Categories. [Conditional_Transfer_Rule] title = Conditional Transfer Rule author = Mihails Milehins topic = Tools date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = This article provides a collection of experimental utilities for unoverloading of definitions and synthesis of conditional transfer rules for the object logic Isabelle/HOL of the formal proof assistant Isabelle written in Isabelle/ML. [Types_To_Sets_Extension] title = Extension of Types-To-Sets author = Mihails Milehins topic = Tools date = 2021-09-06 notify = mihailsmilehins@gmail.com abstract = In their article titled From Types to Sets by Local Type Definitions in Higher-Order Logic and published in the proceedings of the conference Interactive Theorem Proving in 2016, Ondřej Kunčar and Andrei Popescu propose an extension of the logic Isabelle/HOL and an associated algorithm for the relativization of the type-based theorems to more flexible set-based theorems, collectively referred to as Types-To-Sets. One of the aims of their work was to open an opportunity for the development of a software tool for applied relativization in the implementation of the logic Isabelle/HOL of the proof assistant Isabelle. In this article, we provide a prototype of a software framework for the interactive automated relativization of theorems in Isabelle/HOL, developed as an extension of the proof language Isabelle/Isar. The software framework incorporates the implementation of the proposed extension of the logic, and builds upon some of the ideas for further work expressed in the original article on Types-To-Sets by Ondřej Kunčar and Andrei Popescu and the subsequent article Smooth Manifolds and Types to Sets for Linear Algebra in Isabelle/HOL that was written by Fabian Immler and Bohua Zhan and published in the proceedings of the International Conference on Certified Programs and Proofs in 2019. [Complex_Bounded_Operators] title = Complex Bounded Operators author = Jose Manuel Rodriguez Caballero , Dominique Unruh topic = Mathematics/Analysis date = 2021-09-18 notify = unruh@ut.ee abstract = We present a formalization of bounded operators on complex vector spaces. Our formalization contains material on complex vector spaces (normed spaces, Banach spaces, Hilbert spaces) that complements and goes beyond the developments of real vectors spaces in the Isabelle/HOL standard library. We define the type of bounded operators between complex vector spaces (cblinfun) and develop the theory of unitaries, projectors, extension of bounded linear functions (BLT theorem), adjoints, Loewner order, closed subspaces and more. For the finite-dimensional case, we provide code generation support by identifying finite-dimensional operators with matrices as formalized in the Jordan_Normal_Form AFP entry. [Weighted_Path_Order] title = A Formalization of Weighted Path Orders and Recursive Path Orders author = Christian Sternagel , René Thiemann , Akihisa Yamada topic = Logic/Rewriting date = 2021-09-16 notify = rene.thiemann@uibk.ac.at abstract = We define the weighted path order (WPO) and formalize several properties such as strong normalization, the subterm property, and closure properties under substitutions and contexts. Our definition of WPO extends the original definition by also permitting multiset comparisons of arguments instead of just lexicographic extensions. Therefore, our WPO not only subsumes lexicographic path orders (LPO), but also recursive path orders (RPO). We formally prove these subsumptions and therefore all of the mentioned properties of WPO are automatically transferable to LPO and RPO as well. Such a transformation is not required for Knuth–Bendix orders (KBO), since they have already been formalized. Nevertheless, we still provide a proof that WPO subsumes KBO and thereby underline the generality of WPO. [FOL_Axiomatic] title = Soundness and Completeness of an Axiomatic System for First-Order Logic author = Asta Halkjær From topic = Logic/General logic/Classical first-order logic, Logic/Proof theory date = 2021-09-24 notify = ahfrom@dtu.dk abstract = This work is a formalization of the soundness and completeness of an axiomatic system for first-order logic. The proof system is based on System Q1 by Smullyan and the completeness proof follows his textbook "First-Order Logic" (Springer-Verlag 1968). The completeness proof is in the Henkin style where a consistent set is extended to a maximal consistent set using Lindenbaum's construction and Henkin witnesses are added during the construction to ensure saturation as well. The resulting set is a Hintikka set which, by the model existence theorem, is satisfiable in the Herbrand universe. [Virtual_Substitution] title = Verified Quadratic Virtual Substitution for Real Arithmetic author = Matias Scharager , Katherine Cordwell , Stefan Mitsch , André Platzer topic = Computer science/Algorithms/Mathematical date = 2021-10-02 notify = mscharag@cs.cmu.edu, kcordwel@cs.cmu.edu, smitsch@cs.cmu.edu, aplatzer@cs.cmu.edu abstract = This paper presents a formally verified quantifier elimination (QE) algorithm for first-order real arithmetic by linear and quadratic virtual substitution (VS) in Isabelle/HOL. The Tarski-Seidenberg theorem established that the first-order logic of real arithmetic is decidable by QE. However, in practice, QE algorithms are highly complicated and often combine multiple methods for performance. VS is a practically successful method for QE that targets formulas with low-degree polynomials. To our knowledge, this is the first work to formalize VS for quadratic real arithmetic including inequalities. The proofs necessitate various contributions to the existing multivariate polynomial libraries in Isabelle/HOL. Our framework is modularized and easily expandable (to facilitate integrating future optimizations), and could serve as a basis for developing practical general-purpose QE algorithms. Further, as our formalization is designed with practicality in mind, we export our development to SML and test the resulting code on 378 benchmarks from the literature, comparing to Redlog, Z3, Wolfram Engine, and SMT-RAT. This identified inconsistencies in some tools, underscoring the significance of a verified approach for the intricacies of real arithmetic. [Correctness_Algebras] title = Algebras for Iteration, Infinite Executions and Correctness of Sequential Computations author = Walter Guttmann topic = Computer science/Programming languages/Logics date = 2021-10-12 notify = walter.guttmann@canterbury.ac.nz abstract = We study models of state-based non-deterministic sequential computations and describe them using algebras. We propose algebras that describe iteration for strict and non-strict computations. They unify computation models which differ in the fixpoints used to represent iteration. We propose algebras that describe the infinite executions of a computation. They lead to a unified approximation order and results that connect fixpoints in the approximation and refinement orders. This unifies the semantics of recursion for a range of computation models. We propose algebras that describe preconditions and the effect of while-programs under postconditions. They unify correctness statements in two dimensions: one statement applies in various computation models to various correctness claims. [Belief_Revision] title = Belief Revision Theory author = Valentin Fouillard , Safouan Taha , Frédéric Boulanger , Nicolas Sabouret <> topic = Logic/General logic/Logics of knowledge and belief date = 2021-10-19 notify = safouan.taha@lri.fr, valentin.fouillard@limsi.fr abstract = The 1985 paper by Carlos Alchourrón, Peter Gärdenfors, and David Makinson (AGM), “On the Logic of Theory Change: Partial Meet Contraction and Revision Functions” launches a large and rapidly growing literature that employs formal models and logics to handle changing beliefs of a rational agent and to take into account new piece of information observed by this agent. In 2011, a review book titled "AGM 25 Years: Twenty-Five Years of Research in Belief Change" was edited to summarize the first twenty five years of works based on AGM. This HOL-based AFP entry is a faithful formalization of the AGM operators (e.g. contraction, revision, remainder ...) axiomatized in the original paper. It also contains the proofs of all the theorems stated in the paper that show how these operators combine. Both proofs of Harper and Levi identities are established. [X86_Semantics] title = X86 instruction semantics and basic block symbolic execution author = Freek Verbeek , Abhijith Bharadwaj <>, Joshua Bockenek <>, Ian Roessle <>, Timmy Weerwag <>, Binoy Ravindran <> topic = Computer science/Hardware, Computer science/Semantics date = 2021-10-13 notify = freek@vt.edu abstract = This AFP entry provides semantics for roughly 120 different X86-64 assembly instructions. These instructions include various moves, arithmetic/logical operations, jumps, call/return, SIMD extensions and others. External functions are supported by allowing a user to provide custom semantics for these calls. Floating-point operations are mapped to uninterpreted functions. The model provides semantics for register aliasing and a byte-level little-endian memory model. The semantics are purposefully incomplete, but overapproximative. For example, the precise effect of flags may be undefined for certain instructions, or instructions may simply have no semantics at all. In those cases, the semantics are mapped to universally quantified uninterpreted terms from a locale. Second, this entry provides a method to symbolic execution of basic blocks. The method, called ''se_step'' (for: symbolic execution step) fetches an instruction and updates the current symbolic state while keeping track of assumptions made over the memory model. A key component is a set of theorems that prove how reads from memory resolve after writes have occurred. Thirdly, this entry provides a parser that allows the user to copy-paste the output of the standard disassembly tool objdump into Isabelle/HOL. A couple small and explanatory examples are included, including functions from the word count program. Several examples can be supplied upon request (they are not included due to the running time of verification): functions from the floating-point modulo function from FDLIBM, the GLIBC strlen function and the CoreUtils SHA256 implementation. [Registers] title = Quantum and Classical Registers author = Dominique Unruh topic = Computer science/Algorithms/Quantum computing, Computer science/Programming languages/Logics, Computer science/Semantics date = 2021-10-28 notify = unruh@ut.ee -abstract = +abstract = A formalization of the theory of quantum and classical registers as developed by (Unruh, Quantum and Classical Registers). In a nutshell, a register refers to a part of a larger memory or system that can be accessed independently. Registers can be constructed from other registers and several (compatible) registers can be composed. This formalization develops both the generic theory of registers as well as specific instantiations for classical and quantum registers. + +[Szemeredi_Regularity] +title = Szemerédi's Regularity Lemma +author = Chelsea Edmonds , Angeliki Koutsoukou-Argyraki , Lawrence C. Paulson +topic = Mathematics/Graph theory, Mathematics/Combinatorics +date = 2021-11-05 +notify = lp15@cam.ac.uk +abstract = + Szemerédi's + regularity lemma is a key result in the study of large + graphs. It asserts the existence an upper bound on the number of parts + the vertices of a graph need to be partitioned into such that the + edges between the parts are random in a certain sense. This bound + depends only on the desired precision and not on the graph itself, in + the spirit of Ramsey's theorem. The formalisation follows online + course notes by Tim + Gowers and Yufei + Zhao. + +[Factor_Algebraic_Polynomial] +title = Factorization of Polynomials with Algebraic Coefficients +author = Manuel Eberl , René Thiemann +topic = Mathematics/Algebra +date = 2021-11-08 +notify = rene.thiemann@uibk.ac.at +abstract = + The AFP already contains a verified implementation of algebraic + numbers. However, it is has a severe limitation in its factorization + algorithm of real and complex polynomials: the factorization is only + guaranteed to succeed if the coefficients of the polynomial are + rational numbers. In this work, we verify an algorithm to factor all + real and complex polynomials whose coefficients are algebraic. The + existence of such an algorithm proves in a constructive way that the + set of complex algebraic numbers is algebraically closed. Internally, + the algorithm is based on resultants of multivariate polynomials and + an approximation algorithm using interval arithmetic. + +[PAL] +title = Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL +author = Christoph Benzmüller , Sebastian Reiche +topic = Logic/General logic/Logics of knowledge and belief +date = 2021-11-08 +notify = c.benzmueller@gmail.com +abstract = + We present a shallow embedding of public announcement logic (PAL) with + relativized general knowledge in HOL. We then use PAL to obtain an + elegant encoding of the wise men puzzle, which we solve automatically + using sledgehammer. diff --git a/metadata/templates/about.tpl b/metadata/templates/about.tpl --- a/metadata/templates/about.tpl +++ b/metadata/templates/about.tpl @@ -1,78 +1,78 @@ {% extends "base.tpl" %} {% block headline %} Archive of Formal Proofs {% endblock %} {% block content %}

About

The Archive of Formal Proofs is a collection of proof libraries, examples, and larger scientific developments, mechanically checked in the theorem prover Isabelle. It is organized in the way of a scientific journal. Submissions are refereed.

The archive repository is hosted on Heptapod to provide easy free access to archive entries. The entries are tested and maintained continuously against the current stable release of Isabelle. Older versions of archive entries will remain available.

Editors

The editors of the archive are

Why

We aim to strengthen the community and to foster the development of formal proofs.

We hope that the archive will provide

  • a resource of knowledge, examples, and libraries for users,
  • a large and relevant test bed of theories for Isabelle developers, and
  • a central, citable place for authors to publish their theories

We encourage authors of publications that contain Isabelle developments to make their theories available in the Archive of Formal Proofs and to refer to the archive entry in their publication. It makes it easier for referees to check the validity of theorems (all entries in the archive are mechanically checked), it makes it easier for readers of the publication to understand details of your development, and it makes it easier to use and build on your work.

License

All entries in the Archive of Formal Proofs are licensed under a BSD-style License or the GNU LGPL. This means they are free to download, free to use, free to change, and free to redistribute with minimal restrictions.

The authors retain their full copyright on their original work, including their right to make the development available under another, additional license in the future.

{% endblock %} diff --git a/thys/Algebraic_Numbers/Algebraic_Number_Tests.thy b/thys/Algebraic_Numbers/Algebraic_Number_Tests.thy --- a/thys/Algebraic_Numbers/Algebraic_Number_Tests.thy +++ b/thys/Algebraic_Numbers/Algebraic_Number_Tests.thy @@ -1,88 +1,81 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Algebraic Number Tests\ text \We provide a sequence of examples which demonstrate what can be done with the implementation of algebraic numbers.\ theory Algebraic_Number_Tests imports Jordan_Normal_Form.Char_Poly Jordan_Normal_Form.Determinant_Impl Show.Show_Complex "HOL-Library.Code_Target_Nat" "HOL-Library.Code_Target_Int" Berlekamp_Zassenhaus.Factorize_Rat_Poly - Real_Factorization + Complex_Algebraic_Numbers Show_Real_Precise begin subsection \Stand-Alone Examples\ abbreviation (input) "show_lines x \ shows_lines x Nil" fun show_factorization :: "'a :: {semiring_1,show} \ (('a poly \ nat)list) \ string" where "show_factorization (c,[]) = show c" | "show_factorization (c,((p,i) # ps)) = show_factorization (c,ps) @ '' * ('' @ show p @ '')'' @ (if i = 1 then [] else ''^'' @ show i)" definition show_sf_factorization :: "'a :: {semiring_1,show} \ (('a poly \ nat)list) \ string" where "show_sf_factorization x = show_factorization (map_prod id (map (map_prod id Suc)) x) " text \Determine the roots over the rational, real, and complex numbers.\ definition "testpoly = [:5/2, -7/2, 1/2, -5, 7, -1, 5/2, -7/2, 1/2:]" definition "test = show_lines ( real_roots_of_rat_poly testpoly)" value [code] "show_lines ( roots_of_rat_poly testpoly)" value [code] "show_lines ( real_roots_of_rat_poly testpoly)" value [code] "show_lines (complex_roots_of_rat_poly testpoly)" -text \Factorize polynomials over the rational, real, and complex numbers.\ +text \Compute real and complex roots of a polynomial with rational coefficients.\ -value [code] "show_sf_factorization (factorize_rat_poly testpoly)" -value [code] "show_factorization (the (factorize_real_poly testpoly))" -value [code] "show_factorization (the (factorize_complex_poly testpoly))" - -text \If the input is not a rational polynomial, factorization can fail.\ - -value [code] "factorize_real_poly [:sqrt 2,1,3,1:]" text \fails\\\ -value [code] "factorize_real_poly [:sqrt 2,1,3:]" text \does not fail, reveals internal representation\\\ -value [code] "show (factorize_real_poly [:sqrt 2,1,3:])" text \does not fail, pretty printed\ +value [code] "show (complex_roots_of_rat_poly testpoly)" +value [code] "show (real_roots_of_rat_poly testpoly)" text \A sequence of calculations.\ value [code] "show (- sqrt 2 - sqrt 3)" lemma "root 3 4 > sqrt (root 4 3) + \1/10 * root 3 7\" by eval lemma "csqrt (4 + 3 * \) \ \" by eval value [code] "show (csqrt (4 + 3 * \))" value [code] "show (csqrt (1 + \))" subsection \Example Application: Compute Norms of Eigenvalues\ text \For complexity analysis of some matrix $A$ it is important to compute the spectral radius of a matrix, i.e., the maximal norm of all complex eigenvalues, since the spectral radius determines the growth rates of matrix-powers $A^n$, cf.~\cite{JNF-AFP} for a formalized statement of this fact.\ definition eigenvalues :: "rat mat \ complex list" where "eigenvalues A = complex_roots_of_rat_poly (char_poly A)" definition "testmat = mat_of_rows_list 3 [ [1,-4,2], [1/5,7,9], [7,1,5 :: rat] ]" definition "spectral_radius_test = show (Max (set [ norm ev. ev \ eigenvalues testmat]))" value [code] "char_poly testmat" value [code] spectral_radius_test end diff --git a/thys/Algebraic_Numbers/Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Algebraic_Numbers.thy @@ -1,1349 +1,1348 @@ (* Author: René Thiemann Akihisa Yamada Contributors: Manuel Eberl (algebraic integers) License: BSD *) section \Algebraic Numbers: Addition and Multiplication\ text \This theory contains the remaining field operations for algebraic numbers, namely addition and multiplication.\ theory Algebraic_Numbers imports Algebraic_Numbers_Prelim Resultant Polynomial_Factorization.Polynomial_Divisibility begin interpretation coeff_hom: monoid_add_hom "\p. coeff p i" by (unfold_locales, auto) interpretation coeff_hom: comm_monoid_add_hom "\p. coeff p i".. interpretation coeff_hom: group_add_hom "\p. coeff p i".. interpretation coeff_hom: ab_group_add_hom "\p. coeff p i".. interpretation coeff_0_hom: monoid_mult_hom "\p. coeff p 0" by (unfold_locales, auto simp: coeff_mult) interpretation coeff_0_hom: semiring_hom "\p. coeff p 0".. interpretation coeff_0_hom: comm_monoid_mult_hom "\p. coeff p 0".. interpretation coeff_0_hom: comm_semiring_hom "\p. coeff p 0".. subsection \Addition of Algebraic Numbers\ definition "x_y \ [: [: 0, 1 :], -1 :]" definition "poly_x_minus_y p = poly_lift p \\<^sub>p x_y" lemma coeff_xy_power: assumes "k \ n" shows "coeff (x_y ^ n :: 'a :: comm_ring_1 poly poly) k = monom (of_nat (n choose (n - k)) * (- 1) ^ k) (n - k)" proof - define X :: "'a poly poly" where "X = monom (monom 1 1) 0" define Y :: "'a poly poly" where "Y = monom (-1) 1" have [simp]: "monom 1 b * (-1) ^ k = monom ((-1)^k :: 'a) b" for b k by (auto simp: monom_altdef minus_one_power_iff) have "(X + Y) ^ n = (\i\n. of_nat (n choose i) * X ^ i * Y ^ (n - i))" by (subst binomial_ring) auto also have "\ = (\i\n. of_nat (n choose i) * monom (monom ((-1) ^ (n - i)) i) (n - i))" by (simp add: X_def Y_def monom_power mult_monom mult.assoc) also have "\ = (\i\n. monom (monom (of_nat (n choose i) * (-1) ^ (n - i)) i) (n - i))" by (simp add: of_nat_poly smult_monom) also have "coeff \ k = (\i\n. if n - i = k then monom (of_nat (n choose i) * (- 1) ^ (n - i)) i else 0)" by (simp add: of_nat_poly coeff_sum) also have "\ = (\i\{n-k}. monom (of_nat (n choose i) * (- 1) ^ (n - i)) i)" using \k \ n\ by (intro sum.mono_neutral_cong_right) auto also have "X + Y = x_y" by (simp add: X_def Y_def x_y_def monom_altdef) finally show ?thesis using \k \ n\ by simp qed text \The following polynomial represents the sum of two algebraic numbers.\ definition poly_add :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_add p q = resultant (poly_x_minus_y p) (poly_lift q)" subsubsection \@{term poly_add} has desired root\ interpretation poly_x_minus_y_hom: comm_ring_hom poly_x_minus_y by (unfold_locales; simp add: poly_x_minus_y_def hom_distribs) lemma poly2_x_y[simp]: fixes x :: "'a :: comm_ring_1" shows "poly2 x_y x y = x - y" unfolding poly2_def by (simp add: x_y_def) lemma degree_poly_x_minus_y[simp]: fixes p :: "'a::idom poly" shows "degree (poly_x_minus_y p) = degree p" unfolding poly_x_minus_y_def x_y_def by auto lemma poly_x_minus_y_pCons[simp]: "poly_x_minus_y (pCons a p) = [:[: a :]:] + poly_x_minus_y p * x_y" unfolding poly_x_minus_y_def x_y_def by simp lemma poly_poly_poly_x_minus_y[simp]: fixes p :: "'a :: comm_ring_1 poly" shows "poly (poly (poly_x_minus_y p) q) x = poly p (x - poly q x)" by (induct p; simp add: ring_distribs x_y_def) lemma poly2_poly_x_minus_y[simp]: fixes p :: "'a :: comm_ring_1 poly" shows "poly2 (poly_x_minus_y p) x y = poly p (x-y)" unfolding poly2_def by simp interpretation x_y_mult_hom: zero_hom_0 "\p :: 'a :: comm_ring_1 poly poly. x_y * p" proof (unfold_locales) fix p :: "'a poly poly" assume "x_y * p = 0" then show "p = 0" apply (simp add: x_y_def) by (metis eq_neg_iff_add_eq_0 minus_equation_iff minus_pCons synthetic_div_unique_lemma) qed lemma x_y_nonzero[simp]: "x_y \ 0" by (simp add: x_y_def) lemma degree_x_y[simp]: "degree x_y = 1" by (simp add: x_y_def) interpretation x_y_mult_hom: inj_comm_monoid_add_hom "\p :: 'a :: idom poly poly. x_y * p" proof (unfold_locales) show "x_y * p = x_y * q \ p = q" for p q :: "'a poly poly" proof (induct p arbitrary:q) case 0 then show ?case by simp next case p: (pCons a p) from p(3)[unfolded mult_pCons_right] have "x_y * (monom a 0 + pCons 0 1 * p) = x_y * q" apply (subst(asm) pCons_0_as_mult) apply (subst(asm) smult_prod) by (simp only: field_simps distrib_left) then have "monom a 0 + pCons 0 1 * p = q" by simp then show "pCons a p = q" using pCons_as_add by (simp add: monom_0 monom_Suc) qed qed interpretation poly_x_minus_y_hom: inj_idom_hom poly_x_minus_y proof fix p :: "'a poly" assume 0: "poly_x_minus_y p = 0" then have "poly_lift p \\<^sub>p x_y = 0" by (simp add: poly_x_minus_y_def) then show "p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) note p = this[unfolded poly_lift_pCons pcompose_pCons] show ?case proof (cases "a=0") case a0: True with p have "x_y * poly_lift p \\<^sub>p x_y = 0" by simp then have "poly_lift p \\<^sub>p x_y = 0" by simp then show ?thesis using p by simp next case a0: False with p have p0: "p \ 0" by auto from p have "[:[:a:]:] = - x_y * poly_lift p \\<^sub>p x_y" by (simp add: eq_neg_iff_add_eq_0) then have "degree [:[:a:]:] = degree (x_y * poly_lift p \\<^sub>p x_y)" by simp also have "... = degree (x_y::'a poly poly) + degree (poly_lift p \\<^sub>p x_y)" apply (subst degree_mult_eq) apply simp apply (subst pcompose_eq_0) apply (simp add: x_y_def) apply (simp add: p0) apply simp done finally have False by simp then show ?thesis.. qed qed qed lemma poly_add: fixes p q :: "'a ::comm_ring_1 poly" assumes q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" shows "poly (poly_add p q) (x+y) = 0" proof (unfold poly_add_def, rule poly_resultant_zero[OF disjI2]) have "degree q > 0" using poly_zero q0 y by auto thus degq: "degree (poly_lift q) > 0" by auto qed (insert x y, simp_all) subsubsection \@{const poly_add} is nonzero\ text \ We first prove that @{const poly_lift} preserves factorization. The result will be essential also in the next section for division of algebraic numbers. \ interpretation poly_lift_hom: unit_preserving_hom "poly_lift :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly \ _" proof fix x :: "'a poly" assume "poly_lift x dvd 1" then have "poly_y_x (poly_lift x) dvd poly_y_x 1" by simp then show "x dvd 1" by (auto simp add: poly_y_x_poly_lift) qed interpretation poly_lift_hom: factor_preserving_hom "poly_lift::'a::idom poly \ 'a poly poly" proof unfold_locales fix p :: "'a poly" assume p: "irreducible p" show "irreducible (poly_lift p)" proof(rule ccontr) from p have p0: "p \ 0" and "\ p dvd 1" by (auto dest: irreducible_not_unit) with poly_lift_hom.hom_dvd[of p 1] have p1: "\ poly_lift p dvd 1" by auto assume "\ irreducible (poly_lift p)" from this[unfolded irreducible_altdef,simplified] p0 p1 obtain q where "q dvd poly_lift p" and pq: "\ poly_lift p dvd q" and q: "\ q dvd 1" by auto then obtain r where "q * r = poly_lift p" by (elim dvdE, auto) then have "poly_y_x (q * r) = poly_y_x (poly_lift p)" by auto also have "... = [:p:]" by (auto simp: poly_y_x_poly_lift monom_0) also have "poly_y_x (q * r) = poly_y_x q * poly_y_x r" by (auto simp: hom_distribs) finally have "... = [:p:]" by auto then have qp: "poly_y_x q dvd [:p:]" by (metis dvdI) from dvd_const[OF this] p0 have "degree (poly_y_x q) = 0" by auto from degree_0_id[OF this,symmetric] obtain s where qs: "poly_y_x q = [:s:]" by auto have "poly_lift s = poly_y_x (poly_y_x (poly_lift s))" by auto also have "... = poly_y_x [:s:]" by (auto simp: poly_y_x_poly_lift monom_0) also have "... = q" by (auto simp: qs[symmetric]) finally have sq: "poly_lift s = q" by auto from qp[unfolded qs] have sp: "s dvd p" by (auto simp: const_poly_dvd) from irreducibleD'[OF p this] sq q pq show False by auto qed qed text \ We now show that @{const poly_x_minus_y} is a factor-preserving homomorphism. This is essential for this section. This is easy since @{const poly_x_minus_y} can be represented as the composition of two factor-preserving homomorphisms. \ lemma poly_x_minus_y_as_comp: "poly_x_minus_y = (\p. p \\<^sub>p x_y) \ poly_lift" by (intro ext, unfold poly_x_minus_y_def, auto) context idom_isom begin sublocale comm_semiring_isom.. end interpretation poly_x_minus_y_hom: factor_preserving_hom "poly_x_minus_y :: 'a :: idom poly \ 'a poly poly" proof - have \p \\<^sub>p x_y \\<^sub>p x_y = p\ for p :: \'a poly poly\ proof (induction p) case 0 show ?case by simp next case (pCons a p) then show ?case by (unfold x_y_def hom_distribs pcompose_pCons) simp qed then interpret x_y_hom: bijective "\p :: 'a poly poly. p \\<^sub>p x_y" by (unfold bijective_eq_bij) (rule involuntory_imp_bij) interpret x_y_hom: idom_isom "\p :: 'a poly poly. p \\<^sub>p x_y" by standard simp_all have \factor_preserving_hom (\p :: 'a poly poly. p \\<^sub>p x_y)\ and \factor_preserving_hom (poly_lift :: 'a poly \ 'a poly poly)\ .. then show "factor_preserving_hom (poly_x_minus_y :: 'a poly \ _)" by (unfold poly_x_minus_y_as_comp) (rule factor_preserving_hom_comp) qed text \ Now we show that results of @{const poly_x_minus_y} and @{const poly_lift} are coprime. \ lemma poly_y_x_const[simp]: "poly_y_x [:[:a:]:] = [:[:a:]:]" by (simp add: poly_y_x_def monom_0) context begin private abbreviation "y_x == [: [: 0, -1 :], 1 :]" lemma poly_y_x_x_y[simp]: "poly_y_x x_y = y_x" by (simp add: x_y_def poly_y_x_def monom_Suc monom_0) private lemma y_x[simp]: fixes x :: "'a :: comm_ring_1" shows "poly2 y_x x y = y - x" unfolding poly2_def by simp private definition "poly_y_minus_x p \ poly_lift p \\<^sub>p y_x" private lemma poly_y_minus_x_0[simp]: "poly_y_minus_x 0 = 0" by (simp add: poly_y_minus_x_def) private lemma poly_y_minus_x_pCons[simp]: "poly_y_minus_x (pCons a p) = [:[: a :]:] + poly_y_minus_x p * y_x" by (simp add: poly_y_minus_x_def) private lemma poly_y_x_poly_x_minus_y: fixes p :: "'a :: idom poly" shows "poly_y_x (poly_x_minus_y p) = poly_y_minus_x p" apply (induct p, simp) apply (unfold poly_x_minus_y_pCons hom_distribs) by simp lemma degree_poly_y_minus_x[simp]: fixes p :: "'a :: idom poly" shows "degree (poly_y_x (poly_x_minus_y p)) = degree p" by (simp add: poly_y_minus_x_def poly_y_x_poly_x_minus_y) end lemma dvd_all_coeffs_iff: fixes x :: "'a :: comm_semiring_1" (* No addition needed! *) shows "(\pi \ set (coeffs p). x dvd pi) \ (\i. x dvd coeff p i)" (is "?l = ?r") proof- have "?r = (\i\{..degree p} \ {Suc (degree p)..}. x dvd coeff p i)" by auto also have "... = (\i\degree p. x dvd coeff p i)" by (auto simp add: ball_Un coeff_eq_0) also have "... = ?l" by (auto simp: coeffs_def) finally show ?thesis.. qed lemma primitive_imp_no_constant_factor: fixes p :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} poly" assumes pr: "primitive p" and F: "mset_factors F p" and fF: "f \# F" shows "degree f \ 0" proof from F fF have irr: "irreducible f" and fp: "f dvd p" by (auto dest: mset_factors_imp_dvd) assume deg: "degree f = 0" then obtain f0 where f0: "f = [:f0:]" by (auto dest: degree0_coeffs) with fp have "[:f0:] dvd p" by simp then have "f0 dvd coeff p i" for i by (simp add: const_poly_dvd_iff) with primitiveD[OF pr] dvd_all_coeffs_iff have "f0 dvd 1" by (auto simp: coeffs_def) with f0 irr show False by auto qed lemma coprime_poly_x_minus_y_poly_lift: fixes p q :: "'a :: ufd poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and pr: "primitive p" shows "coprime (poly_x_minus_y p) (poly_lift q)" proof(rule ccontr) from degp have p: "\ p dvd 1" by (auto simp: dvd_const) from degp have p0: "p \ 0" by auto from mset_factors_exist[of p, OF p0 p] obtain F where F: "mset_factors F p" by auto with poly_x_minus_y_hom.hom_mset_factors have pF: "mset_factors (image_mset poly_x_minus_y F) (poly_x_minus_y p)" by auto from degq have q: "\ q dvd 1" by (auto simp: dvd_const) from degq have q0: "q \ 0" by auto from mset_factors_exist[OF q0 q] obtain G where G: "mset_factors G q" by auto with poly_lift_hom.hom_mset_factors have pG: "mset_factors (image_mset poly_lift G) (poly_lift q)" by auto assume "\ coprime (poly_x_minus_y p) (poly_lift q)" from this[unfolded not_coprime_iff_common_factor] obtain r where rp: "r dvd (poly_x_minus_y p)" and rq: "r dvd (poly_lift q)" and rU: "\ r dvd 1" by auto note poly_lift_hom.hom_dvd from rp p0 have r0: "r \ 0" by auto from mset_factors_exist[OF r0 rU] obtain H where H: "mset_factors H r" by auto then have "H \ {#}" by auto then obtain h where hH: "h \# H" by fastforce with H mset_factors_imp_dvd have hr: "h dvd r" and h: "irreducible h" by auto from irreducible_not_unit[OF h] have hU: "\ h dvd 1" by auto from hr rp have "h dvd (poly_x_minus_y p)" by (rule dvd_trans) from irreducible_dvd_imp_factor[OF this h pF] p0 obtain f where f: "f \# F" and fh: "poly_x_minus_y f ddvd h" by auto from hr rq have "h dvd (poly_lift q)" by (rule dvd_trans) from irreducible_dvd_imp_factor[OF this h pG] q0 obtain g where g: "g \# G" and gh: "poly_lift g ddvd h" by auto from fh gh have "poly_x_minus_y f ddvd poly_lift g" using ddvd_trans by auto then have "poly_y_x (poly_x_minus_y f) ddvd poly_y_x (poly_lift g)" by simp also have "poly_y_x (poly_lift g) = [:g:]" unfolding poly_y_x_poly_lift monom_0 by auto finally have ddvd: "poly_y_x (poly_x_minus_y f) ddvd [:g:]" by auto then have "degree (poly_y_x (poly_x_minus_y f)) = 0" by (metis degree_pCons_0 dvd_0_left_iff dvd_const) then have "degree f = 0" by simp with primitive_imp_no_constant_factor[OF pr F f] show False by auto qed lemma poly_add_nonzero: fixes p q :: "'a :: ufd poly" assumes p0: "p \ 0" and q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and pr: "primitive p" shows "poly_add p q \ 0" proof have degp: "degree p > 0" using le_0_eq order_degree order_root p0 x by (metis gr0I) have degq: "degree q > 0" using le_0_eq order_degree order_root q0 y by (metis gr0I) assume 0: "poly_add p q = 0" from resultant_zero_imp_common_factor[OF _ this[unfolded poly_add_def]] degp and coprime_poly_x_minus_y_poly_lift[OF degp degq pr] show False by auto qed subsubsection \Summary for addition\ text \Now we lift the results to one that uses @{const ipoly}, by showing some homomorphism lemmas.\ lemma (in comm_ring_hom) map_poly_x_minus_y: "map_poly (map_poly hom) (poly_x_minus_y p) = poly_x_minus_y (map_poly hom p)" proof- interpret mp: map_poly_comm_ring_hom hom.. interpret mmp: map_poly_comm_ring_hom "map_poly hom".. show ?thesis apply (induct p, simp) apply(unfold x_y_def hom_distribs poly_x_minus_y_pCons, simp) done qed lemma (in comm_ring_hom) hom_poly_lift[simp]: "map_poly (map_poly hom) (poly_lift q) = poly_lift (map_poly hom q)" proof - show ?thesis unfolding poly_lift_def unfolding map_poly_map_poly[of coeff_lift,OF coeff_lift_hom.hom_zero] unfolding map_poly_coeff_lift_hom by simp qed lemma lead_coeff_poly_x_minus_y: fixes p :: "'a::idom poly" shows "lead_coeff (poly_x_minus_y p) = [:lead_coeff p * ((- 1) ^ degree p):]" (is "?l = ?r") proof- have "?l = Polynomial.smult (lead_coeff p) ((- 1) ^ degree p)" by (unfold poly_x_minus_y_def, subst lead_coeff_comp; simp add: x_y_def) also have "... = ?r" by (unfold hom_distribs, simp add: smult_as_map_poly[symmetric]) finally show ?thesis. qed lemma degree_coeff_poly_x_minus_y: fixes p q :: "'a :: {idom, semiring_char_0} poly" shows "degree (coeff (poly_x_minus_y p) i) = degree p - i" proof - consider "i = degree p" | "i > degree p" | "i < degree p" by force thus ?thesis proof cases assume "i > degree p" thus ?thesis by (subst coeff_eq_0) auto next assume "i = degree p" thus ?thesis using lead_coeff_poly_x_minus_y[of p] by (simp add: lead_coeff_poly_x_minus_y) next assume "i < degree p" define n where "n = degree p" have "degree (coeff (poly_x_minus_y p) i) = degree (\j\n. [:coeff p j:] * coeff (x_y ^ j) i)" (is "_ = degree (sum ?f _)") by (simp add: poly_x_minus_y_def pcompose_conv_poly poly_altdef coeff_sum n_def) also have "{..n} = insert n {.. = ?f n + sum ?f {.. = n - i" proof - have "degree (?f n) = n - i" using \i < degree p\ by (simp add: n_def coeff_xy_power degree_monom_eq) moreover have "degree (sum ?f {.. {.. j - i" proof (cases "i \ j") case True thus ?thesis by (auto simp: n_def coeff_xy_power degree_monom_eq) next case False hence "coeff (x_y ^ j :: 'a poly poly) i = 0" by (subst coeff_eq_0) (auto simp: degree_power_eq) thus ?thesis by simp qed also have "\ < n - i" using \j \ {.. \i < degree p\ by (auto simp: n_def) finally show "degree ([:coeff p j:] * coeff (x_y ^ j) i) < n - i" . qed (use \i < degree p\ in \auto simp: n_def\) ultimately show ?thesis by (subst degree_add_eq_left) auto qed finally show ?thesis by (simp add: n_def) qed qed lemma coeff_0_poly_x_minus_y [simp]: "coeff (poly_x_minus_y p) 0 = p" by (induction p) (auto simp: poly_x_minus_y_def x_y_def) lemma (in idom_hom) poly_add_hom: assumes p0: "hom (lead_coeff p) \ 0" and q0: "hom (lead_coeff q) \ 0" shows "map_poly hom (poly_add p q) = poly_add (map_poly hom p) (map_poly hom q)" proof - interpret mh: map_poly_idom_hom.. show ?thesis unfolding poly_add_def apply (subst mh.resultant_map_poly(1)[symmetric]) apply (subst degree_map_poly_2) apply (unfold lead_coeff_poly_x_minus_y, unfold hom_distribs, simp add: p0) apply simp apply (subst degree_map_poly_2) apply (simp_all add: q0 map_poly_x_minus_y) done qed lemma(in zero_hom) hom_lead_coeff_nonzero_imp_map_poly_hom: assumes "hom (lead_coeff p) \ 0" shows "map_poly hom p \ 0" proof assume "map_poly hom p = 0" then have "coeff (map_poly hom p) (degree p) = 0" by simp with assms show False by simp qed lemma ipoly_poly_add: fixes x y :: "'a :: idom" assumes p0: "(of_int (lead_coeff p) :: 'a) \ 0" and q0: "(of_int (lead_coeff q) :: 'a) \ 0" and x: "ipoly p x = 0" and y: "ipoly q y = 0" shows "ipoly (poly_add p q) (x+y) = 0" using assms of_int_hom.hom_lead_coeff_nonzero_imp_map_poly_hom[OF q0] by (auto intro: poly_add simp: of_int_hom.poly_add_hom[OF p0 q0]) lemma (in comm_monoid_gcd) gcd_list_eq_0_iff[simp]: "listgcd xs = 0 \ (\x \ set xs. x = 0)" by (induct xs, auto) lemma primitive_field_poly[simp]: "primitive (p :: 'a :: field poly) \ p \ 0" by (unfold primitive_iff_some_content_dvd_1,auto simp: dvd_field_iff coeffs_def) lemma ipoly_poly_add_nonzero: fixes x y :: "'a :: field" assumes "p \ 0" and "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "(of_int (lead_coeff p) :: 'a) \ 0" and "(of_int (lead_coeff q) :: 'a) \ 0" shows "poly_add p q \ 0" proof- from assms have "(of_int_poly (poly_add p q) :: 'a poly) \ 0" apply (subst of_int_hom.poly_add_hom,simp,simp) by (rule poly_add_nonzero, auto dest:of_int_hom.hom_lead_coeff_nonzero_imp_map_poly_hom) then show ?thesis by auto qed lemma represents_add: assumes x: "p represents x" and y: "q represents y" shows "(poly_add p q) represents (x + y)" using assms by (intro representsI ipoly_poly_add ipoly_poly_add_nonzero, auto) subsection \Division of Algebraic Numbers\ definition poly_x_mult_y where [code del]: "poly_x_mult_y p \ (\ i \ degree p. monom (monom (coeff p i) i) i)" lemma coeff_poly_x_mult_y: shows "coeff (poly_x_mult_y p) i = monom (coeff p i) i" (is "?l = ?r") proof(cases "degree p < i") case i: False have "?l = sum (\j. if j = i then (monom (coeff p j) j) else 0) {..degree p}" (is "_ = sum ?f ?A") by (simp add: poly_x_mult_y_def coeff_sum) also have "... = sum ?f {i}" using i by (intro sum.mono_neutral_right, auto) also have "... = ?f i" by simp also have "... = ?r" by auto finally show ?thesis. next case True then show ?thesis by (auto simp: poly_x_mult_y_def coeff_eq_0 coeff_sum) qed lemma poly_x_mult_y_code[code]: "poly_x_mult_y p = (let cs = coeffs p in poly_of_list (map (\ (i, ai). monom ai i) (zip [0 ..< length cs] cs)))" unfolding Let_def poly_of_list_def proof (rule poly_eqI, unfold coeff_poly_x_mult_y) fix n let ?xs = "zip [0.. degree p \ p = 0" unfolding degree_eq_length_coeffs by (cases n, auto) hence "monom (coeff p n) n = 0" using coeff_eq_0[of p n] by auto thus ?thesis unfolding id by simp qed qed definition poly_div :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_div p q = resultant (poly_x_mult_y p) (poly_lift q)" text \@{const poly_div} has desired roots.\ lemma poly2_poly_x_mult_y: fixes p :: "'a :: comm_ring_1 poly" shows "poly2 (poly_x_mult_y p) x y = poly p (x * y)" apply (subst(3) poly_as_sum_of_monoms[symmetric]) apply (unfold poly_x_mult_y_def hom_distribs) by (auto simp: poly2_monom poly_monom power_mult_distrib ac_simps) lemma poly_div: fixes p q :: "'a ::field poly" assumes q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and y0: "y \ 0" shows "poly (poly_div p q) (x/y) = 0" proof (unfold poly_div_def, rule poly_resultant_zero[OF disjI2]) have "degree q > 0" using poly_zero q0 y by auto thus degq: "degree (poly_lift q) > 0" by auto qed (insert x y y0, simp_all add: poly2_poly_x_mult_y) text \@{const poly_div} is nonzero.\ interpretation poly_x_mult_y_hom: ring_hom "poly_x_mult_y :: 'a :: {idom,ring_char_0} poly \ _" by (unfold_locales, auto intro: poly2_ext simp: poly2_poly_x_mult_y hom_distribs) interpretation poly_x_mult_y_hom: inj_ring_hom "poly_x_mult_y :: 'a :: {idom,ring_char_0} poly \ _" proof let ?h = poly_x_mult_y fix f :: "'a poly" assume "?h f = 0" then have "poly2 (?h f) x 1 = 0" for x by simp from this[unfolded poly2_poly_x_mult_y] show "f = 0" by auto qed lemma degree_poly_x_mult_y[simp]: fixes p :: "'a :: {idom, ring_char_0} poly" shows "degree (poly_x_mult_y p) = degree p" (is "?l = ?r") proof(rule antisym) show "?r \ ?l" by (cases "p=0", auto intro: le_degree simp: coeff_poly_x_mult_y) show "?l \ ?r" unfolding poly_x_mult_y_def by (auto intro: degree_sum_le le_trans[OF degree_monom_le]) qed interpretation poly_x_mult_y_hom: unit_preserving_hom "poly_x_mult_y :: 'a :: field_char_0 poly \ _" proof(unfold_locales) let ?h = "poly_x_mult_y :: 'a poly \ _" fix f :: "'a poly" assume unit: "?h f dvd 1" then have "degree (?h f) = 0" and "coeff (?h f) 0 dvd 1" unfolding poly_dvd_1 by auto then have deg: "degree f = 0" by (auto simp add: degree_monom_eq) with unit show "f dvd 1" by(cases "f = 0", auto) qed lemmas poly_y_x_o_poly_lift = o_def[of poly_y_x poly_lift, unfolded poly_y_x_poly_lift] lemma irreducible_dvd_degree: assumes "(f::'a::field poly) dvd g" "irreducible g" "degree f > 0" shows "degree f = degree g" using assms by (metis irreducible_altdef degree_0 dvd_refl is_unit_field_poly linorder_neqE_nat poly_divides_conv0) lemma coprime_poly_x_mult_y_poly_lift: fixes p q :: "'a :: field_char_0 poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and nz: "poly p 0 \ 0 \ poly q 0 \ 0" shows "coprime (poly_x_mult_y p) (poly_lift q)" proof(rule ccontr) from degp have p: "\ p dvd 1" by (auto simp: dvd_const) from degp have p0: "p \ 0" by auto from mset_factors_exist[of p, OF p0 p] obtain F where F: "mset_factors F p" by auto then have pF: "prod_mset (image_mset poly_x_mult_y F) = poly_x_mult_y p" by (auto simp: hom_distribs) from degq have q: "\ is_unit q" by (auto simp: dvd_const) from degq have q0: "q \ 0" by auto from mset_factors_exist[OF q0 q] obtain G where G: "mset_factors G q" by auto with poly_lift_hom.hom_mset_factors have pG: "mset_factors (image_mset poly_lift G) (poly_lift q)" by auto from poly_y_x_hom.hom_mset_factors[OF this] have pG: "mset_factors (image_mset coeff_lift G) [:q:]" by (auto simp: poly_y_x_poly_lift monom_0 image_mset.compositionality poly_y_x_o_poly_lift) assume "\ coprime (poly_x_mult_y p) (poly_lift q)" then have "\ coprime (poly_y_x (poly_x_mult_y p)) (poly_y_x (poly_lift q))" by (simp del: coprime_iff_coprime) from this[unfolded not_coprime_iff_common_factor] obtain r where rp: "r dvd poly_y_x (poly_x_mult_y p)" and rq: "r dvd poly_y_x (poly_lift q)" and rU: "\ r dvd 1" by auto from rp p0 have r0: "r \ 0" by auto from mset_factors_exist[OF r0 rU] obtain H where H: "mset_factors H r" by auto then have "H \ {#}" by auto then obtain h where hH: "h \# H" by fastforce with H mset_factors_imp_dvd have hr: "h dvd r" and h: "irreducible h" by auto from irreducible_not_unit[OF h] have hU: "\ h dvd 1" by auto from hr rp have "h dvd poly_y_x (poly_x_mult_y p)" by (rule dvd_trans) note this[folded pF,unfolded poly_y_x_hom.hom_prod_mset image_mset.compositionality] from prime_elem_dvd_prod_mset[OF h[folded prime_elem_iff_irreducible] this] obtain f where f: "f \# F" and hf: "h dvd poly_y_x (poly_x_mult_y f)" by auto have irrF: "irreducible f" using f F by blast from dvd_trans[OF hr rq] have "h dvd [:q:]" by (simp add: poly_y_x_poly_lift monom_0) from irreducible_dvd_imp_factor[OF this h pG] q0 obtain g where g: "g \# G" and gh: "[:g:] dvd h" by auto from dvd_trans[OF gh hf] have *: "[:g:] dvd poly_y_x (poly_x_mult_y f)" using dvd_trans by auto show False proof (cases "poly f 0 = 0") case f_0: False from poly_hom.hom_dvd[OF *] have "g dvd poly (poly_y_x (poly_x_mult_y f)) [:0:]" by simp also have "... = [:poly f 0:]" by (intro poly_ext, fold poly2_def, simp add: poly2_poly_x_mult_y) also have "... dvd 1" using f_0 by auto finally have "g dvd 1". with g G show False by (auto elim!: mset_factorsE dest!: irreducible_not_unit) next case True hence "[:0,1:] dvd f" by (unfold dvd_iff_poly_eq_0, simp) from irreducible_dvd_degree[OF this irrF] have "degree f = 1" by auto from degree1_coeffs[OF this] True obtain c where c: "c \ 0" and f: "f = [:0,c:]" by auto from g G have irrG: "irreducible g" by auto from poly_hom.hom_dvd[OF *] have "g dvd poly (poly_y_x (poly_x_mult_y f)) 1" by simp also have "\ = f" by (auto simp: f poly_x_mult_y_code Let_def c poly_y_x_pCons map_poly_monom poly_monom poly_lift_def) also have "\ dvd [:0,1:]" unfolding f dvd_def using c by (intro exI[of _ "[: inverse c :]"], auto) finally have g01: "g dvd [:0,1:]" . from divides_degree[OF this] irrG have "degree g = 1" by auto from degree1_coeffs[OF this] obtain a b where g: "g = [:b,a:]" and a: "a \ 0" by auto from g01[unfolded dvd_def] g obtain k where id: "[:0,1:] = g * k" by auto from id have 0: "g \ 0" "k \ 0" by auto from arg_cong[OF id, of degree] have "degree k = 0" unfolding degree_mult_eq[OF 0] unfolding g using a by auto from degree0_coeffs[OF this] obtain kk where k: "k = [:kk:]" by auto from id[unfolded g k] a have "b = 0" by auto hence "poly g 0 = 0" by (auto simp: g) from True this nz \f \# F\ \g \# G\ F G show False by (auto dest!:mset_factors_imp_dvd elim:dvdE) qed qed lemma poly_div_nonzero: fixes p q :: "'a :: field_char_0 poly" assumes p0: "p \ 0" and q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and p_0: "poly p 0 \ 0 \ poly q 0 \ 0" shows "poly_div p q \ 0" proof have degp: "degree p > 0" using le_0_eq order_degree order_root p0 x by (metis gr0I) have degq: "degree q > 0" using le_0_eq order_degree order_root q0 y by (metis gr0I) assume 0: "poly_div p q = 0" from resultant_zero_imp_common_factor[OF _ this[unfolded poly_div_def]] degp and coprime_poly_x_mult_y_poly_lift[OF degp degq] p_0 show False by auto qed subsubsection \Summary for division\ text \Now we lift the results to one that uses @{const ipoly}, by showing some homomorphism lemmas.\ lemma (in inj_comm_ring_hom) poly_x_mult_y_hom: "poly_x_mult_y (map_poly hom p) = map_poly (map_poly hom) (poly_x_mult_y p)" proof - interpret mh: map_poly_inj_comm_ring_hom.. interpret mmh: map_poly_inj_comm_ring_hom "map_poly hom".. show ?thesis unfolding poly_x_mult_y_def by (simp add: hom_distribs) qed lemma (in inj_comm_ring_hom) poly_div_hom: "map_poly hom (poly_div p q) = poly_div (map_poly hom p) (map_poly hom q)" proof - have zero: "\x. hom x = 0 \ x = 0" by simp interpret mh: map_poly_inj_comm_ring_hom.. show ?thesis unfolding poly_div_def mh.resultant_hom[symmetric] by (simp add: poly_x_mult_y_hom) qed lemma ipoly_poly_div: fixes x y :: "'a :: field_char_0" assumes "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "y \ 0" shows "ipoly (poly_div p q) (x/y) = 0" by (unfold of_int_hom.poly_div_hom, rule poly_div, insert assms, auto) lemma ipoly_poly_div_nonzero: fixes x y :: "'a :: field_char_0" assumes "p \ 0" and "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "poly p 0 \ 0 \ poly q 0 \ 0" shows "poly_div p q \ 0" proof- from assms have "(of_int_poly (poly_div p q) :: 'a poly) \ 0" using of_int_hom.poly_map_poly[of p] by (subst of_int_hom.poly_div_hom, subst poly_div_nonzero, auto) then show ?thesis by auto qed lemma represents_div: fixes x y :: "'a :: field_char_0" assumes "p represents x" and "q represents y" and "poly q 0 \ 0" shows "(poly_div p q) represents (x / y)" using assms by (intro representsI ipoly_poly_div ipoly_poly_div_nonzero, auto) subsection \Multiplication of Algebraic Numbers\ definition poly_mult where "poly_mult p q \ poly_div p (reflect_poly q)" lemma represents_mult: assumes px: "p represents x" and qy: "q represents y" and q_0: "poly q 0 \ 0" shows "(poly_mult p q) represents (x * y)" proof- from q_0 qy have y0: "y \ 0" by auto from represents_inverse[OF y0 qy] y0 px q_0 have "poly_mult p q represents x / (inverse y)" unfolding poly_mult_def by (intro represents_div, auto) with y0 show ?thesis by (simp add: field_simps) qed subsection \Summary: Closure Properties of Algebraic Numbers\ lemma algebraic_representsI: "p represents x \ algebraic x" unfolding represents_def algebraic_altdef_ipoly by auto lemma algebraic_of_rat: "algebraic (of_rat x)" by (rule algebraic_representsI[OF poly_rat_represents_of_rat]) lemma algebraic_uminus: "algebraic x \ algebraic (-x)" by (auto dest: algebraic_imp_represents_irreducible intro: algebraic_representsI represents_uminus) lemma algebraic_inverse: "algebraic x \ algebraic (inverse x)" using algebraic_of_rat[of 0] by (cases "x = 0", auto dest: algebraic_imp_represents_irreducible intro: algebraic_representsI represents_inverse) lemma algebraic_plus: "algebraic x \ algebraic y \ algebraic (x + y)" by (auto dest!: algebraic_imp_represents_irreducible_cf_pos intro!: algebraic_representsI[OF represents_add]) lemma algebraic_div: assumes x: "algebraic x" and y: "algebraic y" shows "algebraic (x/y)" proof(cases "y = 0 \ x = 0") case True then show ?thesis using algebraic_of_rat[of 0] by auto next case False then have x0: "x \ 0" and y0: "y \ 0" by auto from x y obtain p q where px: "p represents x" and irr: "irreducible q" and qy: "q represents y" by (auto dest!: algebraic_imp_represents_irreducible) show ?thesis using False px represents_irr_non_0[OF irr qy] by (auto intro!: algebraic_representsI[OF represents_div] qy) qed lemma algebraic_times: "algebraic x \ algebraic y \ algebraic (x * y)" using algebraic_div[OF _ algebraic_inverse, of x y] by (simp add: field_simps) lemma algebraic_root: "algebraic x \ algebraic (root n x)" proof - assume "algebraic x" then obtain p where p: "p represents x" by (auto dest: algebraic_imp_represents_irreducible_cf_pos) from algebraic_representsI[OF represents_nth_root_neg_real[OF _ this, of n]] algebraic_representsI[OF represents_nth_root_pos_real[OF _ this, of n]] algebraic_of_rat[of 0] show ?thesis by (cases "n = 0", force, cases "n > 0", force, cases "n < 0", auto) qed lemma algebraic_nth_root: "n \ 0 \ algebraic x \ y^n = x \ algebraic y" by (auto dest: algebraic_imp_represents_irreducible_cf_pos intro: algebraic_representsI represents_nth_root) - subsection \More on algebraic integers\ (* TODO: this is actually equal to @{term "(-1)^(m*n)"}, but we need a bit more theory on permutations to show this with a reasonable amount of effort. *) definition poly_add_sign :: "nat \ nat \ 'a :: comm_ring_1" where "poly_add_sign m n = signof (\i. if i < n then m + i else if i < m + n then i - n else i)" lemma lead_coeff_poly_add: fixes p q :: "'a :: {idom, semiring_char_0} poly" defines "m \ degree p" and "n \ degree q" assumes "lead_coeff p = 1" "lead_coeff q = 1" "m > 0" "n > 0" shows "lead_coeff (poly_add p q :: 'a poly) = poly_add_sign m n" proof - from assms have [simp]: "p \ 0" "q \ 0" by auto define M where "M = sylvester_mat (poly_x_minus_y p) (poly_lift q)" define \ :: "nat \ nat" where "\ = (\i. if i < n then m + i else if i < m + n then i - n else i)" have \: "\ permutes {0.._def inj_on_def) have nz: "M $$ (i, \ i) \ 0" if "i < m + n" for i using that by (auto simp: M_def \_def sylvester_index_mat m_def n_def) (* have "{(i,j). i \ {.. j \ {.. i < j \ \ i > \ j} = {.. {n.. ?lhs" thus "ij \ ?rhs" by (simp add: \_def split: prod.splits if_splits) auto qed (auto simp: \_def) hence "inversions_on {.. = n * m" by (simp add: inversions_on_def) hence "signof \ = (-1)^(m*n)" using \ by (simp add: signof_def sign_def evenperm_iff_even_inversions) *) have indices_eq: "{0.. (+) n ` {.. \. signof \ * (\i=0.. i)))" have "degree (f \) = degree (\i=0.. i))" using nz by (auto simp: f_def degree_mult_eq sign_def) also have "\ = (\i=0.. i)))" using nz by (subst degree_prod_eq_sum_degree) auto also have "\ = (\i i))) + (\i (n + i))))" by (subst indices_eq, subst sum.union_disjoint) (auto simp: sum.reindex) also have "(\i i))) = (\i_def m_def n_def) also have "(\i (n + i)))) = (\i_def m_def n_def) finally have deg_f1: "degree (f \) = m * n" by simp have deg_f2: "degree (f \) < m * n" if "\ permutes {0.. \ \" for \ proof (cases "\i\{0.. i) = 0") case True hence *: "(\i = 0.. i)) = 0" by auto show ?thesis using \m > 0\ \n > 0\ by (simp add: f_def *) next case False note nz = this from that have \_less: "\ i < m + n" if "i < m + n" for i using permutes_in_image[OF \\ permutes _\] that by auto have "degree (f \) = degree (\i=0.. i))" using nz by (auto simp: f_def degree_mult_eq sign_def) also have "\ = (\i=0.. i)))" using nz by (subst degree_prod_eq_sum_degree) auto also have "\ = (\i i))) + (\i (n + i))))" by (subst indices_eq, subst sum.union_disjoint) (auto simp: sum.reindex) also have "(\i (n + i)))) = (\i_less by (intro sum.cong) (auto simp: M_def sylvester_index_mat \_def m_def n_def) also have "(\i i))) < (\ix\{.. x)) \ m" using \_less by (auto simp: M_def sylvester_index_mat \_def m_def n_def degree_coeff_poly_x_minus_y) next have "\i i \ \ i" proof (rule ccontr) assume nex: "~(\i i \ \ i)" have "\i\m+n-k. \ i = \ i" if "k \ m" for k using that proof (induction k) case 0 thus ?case using \\ permutes _\ \\ permutes _\ by (fastforce simp: permutes_def) next case (Suc k) have IH: "\ i = \ i" if "i \ m+n-k" for i using Suc.prems Suc.IH that by auto from nz have "M $$ (m + n - Suc k, \ (m + n - Suc k)) \ 0" using Suc.prems by auto moreover have "m + n - Suc k \ n" using Suc.prems by auto ultimately have "\ (m+n-Suc k) \ m-Suc k" using assms \_less[of "m+n-Suc k"] Suc.prems by (auto simp: M_def sylvester_index_mat m_def n_def split: if_splits) have "\(\ (m+n-Suc k) > m - Suc k)" proof assume *: "\ (m+n-Suc k) > m - Suc k" have less: "\ (m+n-Suc k) < m" proof (rule ccontr) assume *: "\\ (m + n - Suc k) < m" define j where "j = \ (m + n - Suc k) - m" have "\ (m + n - Suc k) = m + j" using * by (simp add: j_def) moreover { have "j < n" using \_less[of "m+n-Suc k"] \m > 0\ \n > 0\ by (simp add: j_def) hence "\ j = \ j" using nex by auto with \j < n\ have "\ j = m + j" by (auto simp: \_def) } ultimately have "\ (m + n - Suc k) = \ j" by simp hence "m + n - Suc k = j" using permutes_inj[OF \\ permutes _\] unfolding inj_def by blast thus False using \n \ m + n - Suc k\ \_less[of "m+n-Suc k"] \n > 0\ unfolding j_def by linarith qed define j where "j = \ (m+n-Suc k) - (m - Suc k)" from * have j: "\ (m+n-Suc k) = m - Suc k + j" "j > 0" by (auto simp: j_def) have "\ (m+n-Suc k + j) = \ (m+n - Suc k + j)" using * by (intro IH) (auto simp: j_def) also { have "j < Suc k" using less by (auto simp: j_def algebra_simps) hence "m + n - Suc k + j < m + n" using \m > 0\ \n > 0\ Suc.prems by linarith hence "\ (m +n - Suc k + j) = m - Suc k + j" unfolding \_def using Suc.prems by (simp add: \_def) } finally have "\ (m + n - Suc k + j) = \ (m + n - Suc k)" using j by simp hence "m + n - Suc k + j = m + n - Suc k" using permutes_inj[OF \\ permutes _\] unfolding inj_def by blast thus False using \j > 0\ by simp qed with \\ (m+n-Suc k) \ m-Suc k\ have eq: "\ (m+n-Suc k) = m - Suc k" by linarith show ?case proof safe fix i :: nat assume i: "i \ m + n - Suc k" show "\ i = \ i" using eq Suc.prems \m > 0\ IH i proof (cases "i = m + n - Suc k") case True thus ?thesis using eq Suc.prems \m > 0\ by (auto simp: \_def) qed (use IH i in auto) qed qed from this[of m] and nex have "\ i = \ i" for i by (cases "i \ n") auto hence "\ = \" by force thus False using \\ \ \\ by contradiction qed then obtain i where i: "i < n" "\ i \ \ i" by auto have "\ i < m + n" using i by (intro \_less) auto moreover have "\ i = m + i" using i by (auto simp: \_def) ultimately have "degree (M $$ (i, \ i)) < m" using i \m > 0\ by (auto simp: M_def m_def n_def sylvester_index_mat degree_coeff_poly_x_minus_y) thus "\i\{.. i)) < m" using i by blast qed auto finally show "degree (f \) < m * n" by (simp add: mult_ac) qed have "lead_coeff (f \) = poly_add_sign m n" proof - have "lead_coeff (f \) = signof \ * (\i=0.. i)))" by (simp add: f_def sign_def lead_coeff_prod) also have "(\i=0.. i))) = (\i i))) * (\i (n + i))))" by (subst indices_eq, subst prod.union_disjoint) (auto simp: prod.reindex) also have "(\i i))) = (\i_def sylvester_index_mat) also have "(\i (n + i)))) = (\i_def sylvester_index_mat) also have "signof \ = poly_add_sign m n" by (simp add: \_def poly_add_sign_def m_def n_def cong: if_cong) finally show ?thesis using assms by simp qed have "lead_coeff (poly_add p q) = lead_coeff (det (sylvester_mat (poly_x_minus_y p) (poly_lift q)))" by (simp add: poly_add_def resultant_def) also have "det (sylvester_mat (poly_x_minus_y p) (poly_lift q)) = (\\ | \ permutes {0..)" by (simp add: det_def m_def n_def M_def f_def) also have "{\. \ permutes {0.. ({\. \ permutes {0..})" using \ by auto also have "(\\\\. f \) = (\\\{\. \ permutes {0..}. f \) + f \" by (subst sum.insert) (auto simp: finite_permutations) also have "lead_coeff \ = lead_coeff (f \)" proof - have "degree (\\\{\. \ permutes {0..}. f \) < m * n" using assms by (intro degree_sum_smaller deg_f2) (auto simp: m_def n_def finite_permutations) with deg_f1 show ?thesis by (subst lead_coeff_add_le) auto qed finally show ?thesis using \lead_coeff (f \) = _\ by simp qed lemma lead_coeff_poly_mult: fixes p q :: "'a :: {idom, ring_char_0} poly" defines "m \ degree p" and "n \ degree q" assumes "lead_coeff p = 1" "lead_coeff q = 1" "m > 0" "n > 0" assumes "coeff q 0 \ 0" shows "lead_coeff (poly_mult p q :: 'a poly) = 1" proof - from assms have [simp]: "p \ 0" "q \ 0" by auto have [simp]: "degree (reflect_poly q) = n" using assms by (subst degree_reflect_poly_eq) (auto simp: n_def) define M where "M = sylvester_mat (poly_x_mult_y p) (poly_lift (reflect_poly q))" have nz: "M $$ (i, i) \ 0" if "i < m + n" for i using that by (auto simp: M_def sylvester_index_mat m_def n_def coeff_poly_x_mult_y) have indices_eq: "{0.. (+) n ` {.. \. signof \ * (\i=0.. i)))" have "degree (f id) = degree (\i=0.. = (\i=0.. = (\iiiiii) < m * n" if "\ permutes {0.. \ id" for \ proof (cases "\i\{0.. i) = 0") case True hence *: "(\i = 0.. i)) = 0" by auto show ?thesis using \m > 0\ \n > 0\ by (simp add: f_def *) next case False note nz = this from that have \_less: "\ i < m + n" if "i < m + n" for i using permutes_in_image[OF \\ permutes _\] that by auto have "degree (f \) = degree (\i=0.. i))" using nz by (auto simp: f_def degree_mult_eq sign_def) also have "\ = (\i=0.. i)))" using nz by (subst degree_prod_eq_sum_degree) auto also have "\ = (\i i))) + (\i (n + i))))" by (subst indices_eq, subst sum.union_disjoint) (auto simp: sum.reindex) also have "(\i (n + i)))) = (\i_less by (intro sum.cong) (auto simp: M_def sylvester_index_mat m_def n_def) also have "(\i i))) < (\ix\{.. x)) \ m" using \_less by (auto simp: M_def sylvester_index_mat m_def n_def degree_coeff_poly_x_minus_y coeff_poly_x_mult_y intro: order.trans[OF degree_monom_le]) next have "\i i \ i" proof (rule ccontr) assume nex: "\(\i i \ i)" have "\ i = i" for i using that proof (induction i rule: less_induct) case (less i) consider "i < n" | "i \ {n.. m + n" by force thus ?case proof cases assume "i < n" thus ?thesis using nex by auto next assume "i \ m + n" thus ?thesis using \\ permutes _\ by (auto simp: permutes_def) next assume i: "i \ {n.. j = j" if "j < i" for j using that less.prems by (intro less.IH) auto from nz have "M $$ (i, \ i) \ 0" using i by auto hence "\ i \ i" using i \_less[of i] by (auto simp: M_def sylvester_index_mat m_def n_def) moreover have "\ i \ i" proof (rule ccontr) assume *: "\\ i \ i" from * have "\ (\ i) = \ i" by (subst IH) auto hence "\ i = i" using permutes_inj[OF \\ permutes _\] unfolding inj_def by blast with * show False by simp qed ultimately show ?case by simp qed qed hence "\ = id" by force with \\ \ id\ show False by contradiction qed then obtain i where i: "i < n" "\ i \ i" by auto have "\ i < m + n" using i by (intro \_less) auto hence "degree (M $$ (i, \ i)) < m" using i \m > 0\ by (auto simp: M_def m_def n_def sylvester_index_mat degree_coeff_poly_x_minus_y coeff_poly_x_mult_y intro: le_less_trans[OF degree_monom_le]) thus "\i\{.. i)) < m" using i by blast qed auto finally show "degree (f \) < m * n" by (simp add: mult_ac) qed have "lead_coeff (f id) = 1" proof - have "lead_coeff (f id) = (\i=0..i=0..iiiiii\ | \ permutes {0..)" by (simp add: det_def m_def n_def M_def f_def) also have "{\. \ permutes {0... \ permutes {0..\\\. f \) = (\\\{\. \ permutes {0..) + f id" by (subst sum.insert) (auto simp: finite_permutations) also have "lead_coeff \ = lead_coeff (f id)" proof - have "degree (\\\{\. \ permutes {0..) < m * n" using assms by (intro degree_sum_smaller deg_f2) (auto simp: m_def n_def finite_permutations) with deg_f1 show ?thesis by (subst lead_coeff_add_le) auto qed finally show ?thesis using \lead_coeff (f id) = 1\ by simp qed lemma algebraic_int_plus [intro]: fixes x y :: "'a :: field_char_0" assumes "algebraic_int x" "algebraic_int y" shows "algebraic_int (x + y)" proof - from assms(1) obtain p where p: "lead_coeff p = 1" "ipoly p x = 0" by (auto simp: algebraic_int_altdef_ipoly) from assms(2) obtain q where q: "lead_coeff q = 1" "ipoly q y = 0" by (auto simp: algebraic_int_altdef_ipoly) have deg_pos: "degree p > 0" "degree q > 0" using p q by (auto intro!: Nat.gr0I elim!: degree_eq_zeroE) define r where "r = poly_add_sign (degree p) (degree q) * poly_add p q" have "lead_coeff r = 1" using p q deg_pos by (simp add: r_def lead_coeff_mult poly_add_sign_def sign_def lead_coeff_poly_add) moreover have "ipoly r (x + y) = 0" using p q by (simp add: ipoly_poly_add r_def of_int_poly_hom.hom_mult) ultimately show ?thesis by (auto simp: algebraic_int_altdef_ipoly) qed lemma algebraic_int_times [intro]: fixes x y :: "'a :: field_char_0" assumes "algebraic_int x" "algebraic_int y" shows "algebraic_int (x * y)" proof (cases "y = 0") case [simp]: False from assms(1) obtain p where p: "lead_coeff p = 1" "ipoly p x = 0" by (auto simp: algebraic_int_altdef_ipoly) from assms(2) obtain q where q: "lead_coeff q = 1" "ipoly q y = 0" by (auto simp: algebraic_int_altdef_ipoly) have deg_pos: "degree p > 0" "degree q > 0" using p q by (auto intro!: Nat.gr0I elim!: degree_eq_zeroE) have [simp]: "q \ 0" using q by auto define n where "n = Polynomial.order 0 q" have "monom 1 n dvd q" by (simp add: n_def monom_1_dvd_iff) then obtain q' where q_split: "q = q' * monom 1 n" by auto have "Polynomial.order 0 q = Polynomial.order 0 q' + n" using \q \ 0\ unfolding q_split by (subst order_mult) auto hence "poly q' 0 \ 0" unfolding n_def using \q \ 0\ by (simp add: q_split order_root) have q': "ipoly q' y = 0" "lead_coeff q' = 1" using q_split q by (auto simp: of_int_poly_hom.hom_mult poly_monom lead_coeff_mult degree_monom_eq) from this have deg_pos': "degree q' > 0" by (intro Nat.gr0I) (auto elim!: degree_eq_zeroE) from \poly q' 0 \ 0\ have [simp]: "coeff q' 0 \ 0" by (auto simp: monom_1_dvd_iff' poly_0_coeff_0) have "p represents x" "q' represents y" using p q' by (auto simp: represents_def) hence "poly_mult p q' represents x * y" by (rule represents_mult) (simp add: poly_0_coeff_0) moreover have "lead_coeff (poly_mult p q') = 1" using p deg_pos q' deg_pos' by (simp add: lead_coeff_mult lead_coeff_poly_mult) ultimately show ?thesis by (auto simp: algebraic_int_altdef_ipoly represents_def) qed auto lemma algebraic_int_power [intro]: "algebraic_int (x :: 'a :: field_char_0) \ algebraic_int (x ^ n)" by (induction n) auto lemma algebraic_int_diff [intro]: fixes x y :: "'a :: field_char_0" assumes "algebraic_int x" "algebraic_int y" shows "algebraic_int (x - y)" using algebraic_int_plus[OF assms(1) algebraic_int_minus[OF assms(2)]] by simp lemma algebraic_int_sum [intro]: "(\x. x \ A \ algebraic_int (f x :: 'a :: field_char_0)) \ algebraic_int (sum f A)" by (induction A rule: infinite_finite_induct) auto lemma algebraic_int_prod [intro]: "(\x. x \ A \ algebraic_int (f x :: 'a :: field_char_0)) \ algebraic_int (prod f A)" by (induction A rule: infinite_finite_induct) auto lemma algebraic_int_nth_root_real_iff: "algebraic_int (root n x) \ n = 0 \ algebraic_int x" proof - have "algebraic_int x" if "algebraic_int (root n x)" "n \ 0" proof - from that(1) have "algebraic_int (root n x ^ n)" by auto also have "root n x ^ n = (if even n then \x\ else x)" using sgn_power_root[of n x] that(2) by (auto simp: sgn_if split: if_splits) finally show ?thesis by (auto split: if_splits) qed thus ?thesis by auto qed lemma algebraic_int_power_iff: "algebraic_int (x ^ n :: 'a :: field_char_0) \ n = 0 \ algebraic_int x" proof - have "algebraic_int x" if "algebraic_int (x ^ n)" "n > 0" proof (rule algebraic_int_root) show "poly (monom 1 n) x = x ^ n" by (auto simp: poly_monom) qed (use that in \auto simp: degree_monom_eq\) thus ?thesis by auto qed lemma algebraic_int_power_iff' [simp]: "n > 0 \ algebraic_int (x ^ n :: 'a :: field_char_0) \ algebraic_int x" by (subst algebraic_int_power_iff) auto lemma algebraic_int_sqrt_iff [simp]: "algebraic_int (sqrt x) \ algebraic_int x" by (simp add: sqrt_def algebraic_int_nth_root_real_iff) lemma algebraic_int_csqrt_iff [simp]: "algebraic_int (csqrt x) \ algebraic_int x" proof assume "algebraic_int (csqrt x)" hence "algebraic_int (csqrt x ^ 2)" by (rule algebraic_int_power) thus "algebraic_int x" by simp qed auto lemma algebraic_int_norm_complex [intro]: assumes "algebraic_int (z :: complex)" shows "algebraic_int (norm z)" proof - from assms have "algebraic_int (z * cnj z)" by auto also have "z * cnj z = of_real (norm z ^ 2)" by (rule complex_norm_square [symmetric]) finally show ?thesis by simp qed hide_const (open) x_y end diff --git a/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy b/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy --- a/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy +++ b/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy @@ -1,1082 +1,1091 @@ (* Author: René Thiemann Sebastiaan Joosten Akihisa Yamada License: BSD *) section \Algebraic Numbers -- Excluding Addition and Multiplication\ text \This theory contains basic definition and results on algebraic numbers, namely that algebraic numbers are closed under negation, inversion, $n$-th roots, and that every rational number is algebraic. For all of these closure properties, corresponding polynomial witnesses are available. Moreover, this theory contains the uniqueness result, that for every algebraic number there is exactly one content-free irreducible polynomial with positive leading coefficient for it. This result is stronger than similar ones which you find in many textbooks. The reason is that here we do not require a least degree construction. This is essential, since given some content-free irreducible polynomial for x, how should we check whether the degree is optimal. In the formalized result, this is not required. The result is proven via GCDs, and that the GCD does not change when executed on the rational numbers or on the reals or complex numbers, and that the GCD of a rational polynomial can be expressed via the GCD of integer polynomials.\ text \Many results are taken from the textbook \cite[pages 317ff]{AlgNumbers}.\ theory Algebraic_Numbers_Prelim imports "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" Polynomial_Interpolation.Newton_Interpolation (* for lemma smult_1 *) Polynomial_Factorization.Gauss_Lemma Berlekamp_Zassenhaus.Unique_Factorization_Poly Polynomial_Factorization.Square_Free_Factorization begin lemma primitive_imp_unit_iff: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes pr: "primitive p" shows "p dvd 1 \ degree p = 0" proof assume "degree p = 0" from degree0_coeffs[OF this] obtain p0 where p: "p = [:p0:]" by auto then have "\c \ set (coeffs p). p0 dvd c" by (simp add: cCons_def) with pr have "p0 dvd 1" by (auto dest: primitiveD) with p show "p dvd 1" by auto next assume "p dvd 1" then show "degree p = 0" by (auto simp: poly_dvd_1) qed lemma dvd_all_coeffs_imp_dvd: assumes "\a \ set (coeffs p). c dvd a" shows "[:c:] dvd p" proof (insert assms, induct p) case 0 then show ?case by simp next case (pCons a p) have "pCons a p = [:a:] + pCons 0 p" by simp also have "[:c:] dvd ..." proof (rule dvd_add) from pCons show "[:c:] dvd [:a:]" by (auto simp: cCons_def) from pCons have "[:c:] dvd p" by auto from Rings.dvd_mult[OF this] show "[:c:] dvd pCons 0 p" by (subst pCons_0_as_mult) qed finally show ?case. qed lemma irreducible_content: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "irreducible p" shows "degree p = 0 \ primitive p" proof(rule ccontr) assume not: "\?thesis" then obtain c where c1: "\c dvd 1" and "\a \ set (coeffs p). c dvd a" by (auto elim: not_primitiveE) from dvd_all_coeffs_imp_dvd[OF this(2)] obtain r where p: "p = r * [:c:]" by (elim dvdE, auto) from irreducibleD[OF assms this] have "r dvd 1 \ [:c:] dvd 1" by auto with c1 have "r dvd 1" unfolding const_poly_dvd_1 by auto then have "degree r = 0" unfolding poly_dvd_1 by auto with p have "degree p = 0" by auto with not show False by auto qed (* TODO: move *) lemma linear_irreducible_field: fixes p :: "'a :: field poly" assumes deg: "degree p = 1" shows "irreducible p" proof (intro irreducibleI) from deg show p0: "p \ 0" by auto from deg show "\ p dvd 1" by (auto simp: poly_dvd_1) fix a b assume p: "p = a * b" with p0 have a0: "a \ 0" and b0: "b \ 0" by auto from degree_mult_eq[OF this, folded p] assms consider "degree a = 1" "degree b = 0" | "degree a = 0" "degree b = 1" by force then show "a dvd 1 \ b dvd 1" by (cases; insert a0 b0, auto simp: primitive_imp_unit_iff) qed (* TODO: move *) lemma linear_irreducible_int: fixes p :: "int poly" assumes deg: "degree p = 1" and cp: "content p dvd 1" shows "irreducible p" proof (intro irreducibleI) from deg show p0: "p \ 0" by auto from deg show "\ p dvd 1" by (auto simp: poly_dvd_1) fix a b assume p: "p = a * b" note * = cp[unfolded p is_unit_content_iff, unfolded content_mult] have a1: "content a dvd 1" and b1: "content b dvd 1" using content_ge_0_int[of a] pos_zmult_eq_1_iff_lemma[OF *] * by (auto simp: abs_mult) with p0 have a0: "a \ 0" and b0: "b \ 0" by auto from degree_mult_eq[OF this, folded p] assms consider "degree a = 1" "degree b = 0" | "degree a = 0" "degree b = 1" by force then show "a dvd 1 \ b dvd 1" by (cases; insert a1 b1, auto simp: primitive_imp_unit_iff) qed lemma irreducible_connect_rev: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes irr: "irreducible p" and deg: "degree p > 0" shows "irreducible\<^sub>d p" proof(intro irreducible\<^sub>dI deg) fix q r assume degq: "degree q > 0" and diff: "degree q < degree p" and p: "p = q * r" from degq have nu: "\ q dvd 1" by (auto simp: poly_dvd_1) from irreducibleD[OF irr p] nu have "r dvd 1" by auto then have "degree r = 0" by (auto simp: poly_dvd_1) with degq diff show False unfolding p using degree_mult_le[of q r] by auto qed subsection \Polynomial Evaluation of Integer and Rational Polynomials in Fields.\ abbreviation ipoly where "ipoly f x \ poly (of_int_poly f) x" lemma poly_map_poly_code[code_unfold]: "poly (map_poly h p) x = fold_coeffs (\ a b. h a + x * b) p 0" by (induct p, auto) abbreviation real_of_int_poly :: "int poly \ real poly" where "real_of_int_poly \ of_int_poly" abbreviation real_of_rat_poly :: "rat poly \ real poly" where "real_of_rat_poly \ map_poly of_rat" lemma of_rat_of_int[simp]: "of_rat \ of_int = of_int" by auto lemma ipoly_of_rat[simp]: "ipoly p (of_rat y) = of_rat (ipoly p y)" proof- have id: "of_int = of_rat o of_int" unfolding comp_def by auto show ?thesis by (subst id, subst map_poly_map_poly[symmetric], auto) qed lemma ipoly_of_real[simp]: "ipoly p (of_real x :: 'a :: {field,real_algebra_1}) = of_real (ipoly p x)" proof - have id: "of_int = of_real o of_int" unfolding comp_def by auto show ?thesis by (subst id, subst map_poly_map_poly[symmetric], auto) qed lemma finite_ipoly_roots: assumes "p \ 0" shows "finite {x :: real. ipoly p x = 0}" proof - let ?p = "real_of_int_poly p" from assms have "?p \ 0" by auto thus ?thesis by (rule poly_roots_finite) qed subsection \Algebraic Numbers -- Definition, Inverse, and Roots\ text \A number @{term "x :: 'a :: field"} is algebraic iff it is the root of an integer polynomial. Whereas the Isabelle distribution this is defined via the embedding of integers in an field via @{const Ints}, we work with integer polynomials of type @{type int} and then use @{const ipoly} for evaluating the polynomial at a real or complex point.\ lemma algebraic_altdef_ipoly: shows "algebraic x \ (\p. ipoly p x = 0 \ p \ 0)" unfolding algebraic_def proof (safe, goal_cases) case (1 p) define the_int where "the_int = (\x::'a. THE r. x = of_int r)" define p' where "p' = map_poly the_int p" have of_int_the_int: "of_int (the_int x) = x" if "x \ \" for x unfolding the_int_def by (rule sym, rule theI') (insert that, auto simp: Ints_def) have the_int_0_iff: "the_int x = 0 \ x = 0" if "x \ \" using of_int_the_int[OF that] by auto have "map_poly of_int p' = map_poly (of_int \ the_int) p" by (simp add: p'_def map_poly_map_poly) also from 1 of_int_the_int have "\ = p" by (subst poly_eq_iff) (auto simp: coeff_map_poly) finally have p_p': "map_poly of_int p' = p" . show ?case proof (intro exI conjI notI) from 1 show "ipoly p' x = 0" by (simp add: p_p') next assume "p' = 0" hence "p = 0" by (simp add: p_p' [symmetric]) with \p \ 0\ show False by contradiction qed next case (2 p) thus ?case by (intro exI[of _ "map_poly of_int p"], auto) qed text \Definition of being algebraic with explicit witness polynomial.\ definition represents :: "int poly \ 'a :: field_char_0 \ bool" (infix "represents" 51) where "p represents x = (ipoly p x = 0 \ p \ 0)" lemma representsI[intro]: "ipoly p x = 0 \ p \ 0 \ p represents x" unfolding represents_def by auto lemma representsD: assumes "p represents x" shows "p \ 0" and "ipoly p x = 0" using assms unfolding represents_def by auto lemma representsE: assumes "p represents x" and "p \ 0 \ ipoly p x = 0 \ thesis" shows thesis using assms unfolding represents_def by auto lemma represents_imp_degree: fixes x :: "'a :: field_char_0" assumes "p represents x" shows "degree p \ 0" proof- from assms have "p \ 0" and px: "ipoly p x = 0" by (auto dest:representsD) then have "(of_int_poly p :: 'a poly) \ 0" by auto then have "degree (of_int_poly p :: 'a poly) \ 0" by (fold poly_zero[OF px]) then show ?thesis by auto qed lemma representsE_full[elim]: assumes rep: "p represents x" and main: "p \ 0 \ ipoly p x = 0 \ degree p \ 0 \ thesis" shows thesis by (rule main, insert represents_imp_degree[OF rep] rep, auto elim: representsE) lemma represents_of_rat[simp]: "p represents (of_rat x) = p represents x" by (auto elim!:representsE) lemma represents_of_real[simp]: "p represents (of_real x) = p represents x" by (auto elim!:representsE) lemma algebraic_iff_represents: "algebraic x \ (\ p. p represents x)" unfolding algebraic_altdef_ipoly represents_def .. lemma represents_irr_non_0: assumes irr: "irreducible p" and ap: "p represents x" and x0: "x \ 0" shows "poly p 0 \ 0" proof have nu: "\ [:0,1::int:] dvd 1" by (auto simp: poly_dvd_1) assume "poly p 0 = 0" hence dvd: "[: 0, 1 :] dvd p" by (unfold dvd_iff_poly_eq_0, simp) then obtain q where pq: "p = [:0,1:] * q" by (elim dvdE) from irreducibleD[OF irr this] nu have "q dvd 1" by auto from this obtain r where "q = [:r:]" "r dvd 1" by (auto simp add: poly_dvd_1 dest: degree0_coeffs) with pq have "p = [:0,r:]" by auto with ap have "x = 0" by (auto simp: of_int_hom.map_poly_pCons_hom) with x0 show False by auto qed text \The polynomial encoding a rational number.\ definition poly_rat :: "rat \ int poly" where "poly_rat x = (case quotient_of x of (n,d) \ [:-n,d:])" definition abs_int_poly:: "int poly \ int poly" where "abs_int_poly p \ if lead_coeff p < 0 then -p else p" lemma pos_poly_abs_poly[simp]: shows "lead_coeff (abs_int_poly p) > 0 \ p \ 0" proof- have "p \ 0 \ lead_coeff p * sgn (lead_coeff p) > 0" by (fold abs_sgn, auto) then show ?thesis by (auto simp: abs_int_poly_def mult.commute) qed lemma abs_int_poly_0[simp]: "abs_int_poly 0 = 0" by (auto simp: abs_int_poly_def) lemma abs_int_poly_eq_0_iff[simp]: "abs_int_poly p = 0 \ p = 0" by (auto simp: abs_int_poly_def sgn_eq_0_iff) lemma degree_abs_int_poly[simp]: "degree (abs_int_poly p) = degree p" by (auto simp: abs_int_poly_def sgn_eq_0_iff) lemma abs_int_poly_dvd[simp]: "abs_int_poly p dvd q \ p dvd q" by (unfold abs_int_poly_def, auto) (*TODO: move & generalize *) lemma (in idom) irreducible_uminus[simp]: "irreducible (-x) \ irreducible x" proof- have "-x = -1 * x" by simp also have "irreducible ... \ irreducible x" by (rule irreducible_mult_unit_left, auto) finally show ?thesis. qed lemma irreducible_abs_int_poly[simp]: "irreducible (abs_int_poly p) \ irreducible p" by (unfold abs_int_poly_def, auto) lemma coeff_abs_int_poly[simp]: "coeff (abs_int_poly p) n = (if lead_coeff p < 0 then - coeff p n else coeff p n)" by (simp add: abs_int_poly_def) lemma lead_coeff_abs_int_poly[simp]: "lead_coeff (abs_int_poly p) = abs (lead_coeff p)" by auto lemma ipoly_abs_int_poly_eq_zero_iff[simp]: "ipoly (abs_int_poly p) (x :: 'a :: comm_ring_1) = 0 \ ipoly p x = 0" by (auto simp: abs_int_poly_def sgn_eq_0_iff of_int_poly_hom.hom_uminus) lemma abs_int_poly_represents[simp]: "abs_int_poly p represents x \ p represents x" by (auto elim!:representsE) (* TODO: Move *) lemma content_pCons[simp]: "content (pCons a p) = gcd a (content p)" by (unfold content_def coeffs_pCons_eq_cCons cCons_def, auto) lemma content_uminus[simp]: fixes p :: "'a :: ring_gcd poly" shows "content (-p) = content p" by (induct p, auto) lemma primitive_abs_int_poly[simp]: "primitive (abs_int_poly p) \ primitive p" by (auto simp: abs_int_poly_def) lemma abs_int_poly_inv[simp]: "smult (sgn (lead_coeff p)) (abs_int_poly p) = p" by (cases "lead_coeff p > 0", auto simp: abs_int_poly_def) definition cf_pos :: "int poly \ bool" where "cf_pos p = (content p = 1 \ lead_coeff p > 0)" definition cf_pos_poly :: "int poly \ int poly" where "cf_pos_poly f = (let c = content f; d = (sgn (lead_coeff f) * c) in sdiv_poly f d)" lemma sgn_is_unit[intro!]: fixes x :: "'a :: linordered_idom" (* find/make better class *) assumes "x \ 0" shows "sgn x dvd 1" using assms by(cases x "0::'a" rule:linorder_cases, auto) lemma cf_pos_poly_0[simp]: "cf_pos_poly 0 = 0" by (unfold cf_pos_poly_def sdiv_poly_def, auto) lemma cf_pos_poly_eq_0[simp]: "cf_pos_poly f = 0 \ f = 0" proof(cases "f = 0") case True thus ?thesis unfolding cf_pos_poly_def Let_def by (simp add: sdiv_poly_def) next case False then have lc0: "lead_coeff f \ 0" by auto then have s0: "sgn (lead_coeff f) \ 0" (is "?s \ 0") and "content f \ 0" (is "?c \ 0") by (auto simp: sgn_0_0) then have sc0: "?s * ?c \ 0" by auto { fix i from content_dvd_coeff sgn_is_unit[OF lc0] have "?s * ?c dvd coeff f i" by (auto simp: unit_dvd_iff) then have "coeff f i div (?s * ?c) = 0 \ coeff f i = 0" by (auto simp:dvd_div_eq_0_iff) } note * = this show ?thesis unfolding cf_pos_poly_def Let_def sdiv_poly_def poly_eq_iff by (auto simp: coeff_map_poly *) qed lemma shows cf_pos_poly_main: "smult (sgn (lead_coeff f) * content f) (cf_pos_poly f) = f" (is ?g1) and content_cf_pos_poly[simp]: "content (cf_pos_poly f) = (if f = 0 then 0 else 1)" (is ?g2) and lead_coeff_cf_pos_poly[simp]: "lead_coeff (cf_pos_poly f) > 0 \ f \ 0" (is ?g3) and cf_pos_poly_dvd[simp]: "cf_pos_poly f dvd f" (is ?g4) proof(atomize(full), (cases "f = 0"; intro conjI)) case True then show ?g1 ?g2 ?g3 ?g4 by simp_all next case f0: False let ?s = "sgn (lead_coeff f)" have s: "?s \ {-1,1}" using f0 unfolding sgn_if by auto define g where "g \ smult ?s f" define d where "d \ ?s * content f" have "content g = content ([:?s:] * f)" unfolding g_def by simp also have "\ = content [:?s:] * content f" unfolding content_mult by simp also have "content [:?s:] = 1" using s by (auto simp: content_def) finally have cg: "content g = content f" by simp from f0 have d: "cf_pos_poly f = sdiv_poly f d" by (auto simp: cf_pos_poly_def Let_def d_def) let ?g = "primitive_part g" define ng where "ng = primitive_part g" note d also have "sdiv_poly f d = sdiv_poly g (content g)" unfolding cg unfolding g_def d_def by (rule poly_eqI, unfold coeff_sdiv_poly coeff_smult, insert s, auto simp: div_minus_right) finally have fg: "cf_pos_poly f = primitive_part g" unfolding primitive_part_alt_def . have "lead_coeff f \ 0" using f0 by auto hence lg: "lead_coeff g > 0" unfolding g_def lead_coeff_smult by (meson linorder_neqE_linordered_idom sgn_greater sgn_less zero_less_mult_iff) hence g0: "g \ 0" by auto from f0 content_primitive_part[OF this] show ?g2 unfolding fg by auto from g0 have "content g \ 0" by simp with arg_cong[OF content_times_primitive_part[of g], of lead_coeff, unfolded lead_coeff_smult] lg content_ge_0_int[of g] have lg': "lead_coeff ng > 0" unfolding ng_def by (metis dual_order.antisym dual_order.strict_implies_order zero_less_mult_iff) with f0 show ?g3 unfolding fg ng_def by auto have d0: "d \ 0" using s f0 by (force simp add: d_def) have "smult d (cf_pos_poly f) = smult ?s (smult (content f) (sdiv_poly (smult ?s f) (content f)))" unfolding fg primitive_part_alt_def cg by (simp add: g_def d_def) also have "sdiv_poly (smult ?s f) (content f) = smult ?s (sdiv_poly f (content f))" using s by (metis cg g_def primitive_part_alt_def primitive_part_smult_int sgn_sgn) finally have "smult d (cf_pos_poly f) = smult (content f) (primitive_part f)" unfolding primitive_part_alt_def using s by auto also have "\ = f" by (rule content_times_primitive_part) finally have df: "smult d (cf_pos_poly f) = f" . with d0 show ?g1 by (auto simp: d_def) from df have *: "f = cf_pos_poly f * [:d:]" by simp from dvdI[OF this] show ?g4. qed (* TODO: remove *) lemma irreducible_connect_int: fixes p :: "int poly" assumes ir: "irreducible\<^sub>d p" and c: "content p = 1" shows "irreducible p" using c primitive_iff_content_eq_1 ir irreducible_primitive_connect by blast lemma fixes x :: "'a :: {idom,ring_char_0}" shows ipoly_cf_pos_poly_eq_0[simp]: "ipoly (cf_pos_poly p) x = 0 \ ipoly p x = 0" and degree_cf_pos_poly[simp]: "degree (cf_pos_poly p) = degree p" and cf_pos_cf_pos_poly[intro]: "p \ 0 \ cf_pos (cf_pos_poly p)" proof- show "degree (cf_pos_poly p) = degree p" by (subst(3) cf_pos_poly_main[symmetric], auto simp:sgn_eq_0_iff) { assume p: "p \ 0" show "cf_pos (cf_pos_poly p)" using cf_pos_poly_main p by (auto simp: cf_pos_def) have "(ipoly (cf_pos_poly p) x = 0) = (ipoly p x = 0)" apply (subst(3) cf_pos_poly_main[symmetric]) by (auto simp: sgn_eq_0_iff hom_distribs) } then show "(ipoly (cf_pos_poly p) x = 0) = (ipoly p x = 0)" by (cases "p = 0", auto) qed lemma cf_pos_poly_eq_1: "cf_pos_poly f = 1 \ degree f = 0 \ f \ 0" (is "?l \ ?r") proof(intro iffI conjI) assume ?r then have df0: "degree f = 0" and f0: "f \ 0" by auto from degree0_coeffs[OF df0] obtain f0 where f: "f = [:f0:]" by auto show "cf_pos_poly f = 1" using f0 unfolding f cf_pos_poly_def Let_def sdiv_poly_def by (auto simp: content_def mult_sgn_abs) next assume l: ?l then have "degree (cf_pos_poly f) = 0" by auto then show "degree f = 0" by simp from l have "cf_pos_poly f \ 0" by auto then show "f \ 0" by simp qed lemma irr_cf_poly_rat[simp]: "irreducible (poly_rat x)" "lead_coeff (poly_rat x) > 0" "primitive (poly_rat x)" proof - obtain n d where x: "quotient_of x = (n,d)" by force hence id: "poly_rat x = [:-n,d:]" by (auto simp: poly_rat_def) from quotient_of_denom_pos[OF x] have d: "d > 0" by auto show "lead_coeff (poly_rat x) > 0" "primitive (poly_rat x)" unfolding id cf_pos_def using d quotient_of_coprime[OF x] by (auto simp: content_def) from this[unfolded cf_pos_def] show irr: "irreducible (poly_rat x)" unfolding id using d by (auto intro!: linear_irreducible_int) qed lemma poly_rat[simp]: "ipoly (poly_rat x) (of_rat x :: 'a :: field_char_0) = 0" "ipoly (poly_rat x) x = 0" "poly_rat x \ 0" "ipoly (poly_rat x) y = 0 \ y = (of_rat x :: 'a)" proof - from irr_cf_poly_rat(1)[of x] show "poly_rat x \ 0" unfolding Factorial_Ring.irreducible_def by auto obtain n d where x: "quotient_of x = (n,d)" by force hence id: "poly_rat x = [:-n,d:]" by (auto simp: poly_rat_def) from quotient_of_denom_pos[OF x] have d: "d \ 0" by auto have "y * of_int d = of_int n \ y = of_int n / of_int d" using d by (simp add: eq_divide_imp) with d id show "ipoly (poly_rat x) (of_rat x) = 0" "ipoly (poly_rat x) x = 0" "ipoly (poly_rat x) y = 0 \ y = (of_rat x :: 'a)" by (auto simp: of_rat_minus of_rat_divide simp: quotient_of_div[OF x]) qed lemma poly_rat_represents_of_rat: "(poly_rat x) represents (of_rat x)" by auto lemma ipoly_smult_0_iff: assumes c: "c \ 0" shows "(ipoly (smult c p) x = (0 :: real)) = (ipoly p x = 0)" using c by (simp add: hom_distribs) (* TODO *) lemma not_irreducibleD: assumes "\ irreducible x" and "x \ 0" and "\ x dvd 1" shows "\y z. x = y * z \ \ y dvd 1 \ \ z dvd 1" using assms apply (unfold Factorial_Ring.irreducible_def) by auto lemma cf_pos_poly_represents[simp]: "(cf_pos_poly p) represents x \ p represents x" unfolding represents_def by auto lemma coprime_prod: (* TODO: move *) "a \ 0 \ c \ 0 \ coprime (a * b) (c * d) \ coprime b (d::'a::{semiring_gcd})" by auto lemma smult_prod: (* TODO: move or find corresponding lemma *) "smult a b = monom a 0 * b" by (simp add: monom_0) lemma degree_map_poly_2: assumes "f (lead_coeff p) \ 0" shows "degree (map_poly f p) = degree p" proof (cases "p=0") case False thus ?thesis unfolding degree_eq_length_coeffs Polynomial.coeffs_map_poly using assms by (simp add:coeffs_def) qed auto lemma irreducible_cf_pos_poly: assumes irr: "irreducible p" and deg: "degree p \ 0" shows "irreducible (cf_pos_poly p)" (is "irreducible ?p") proof (unfold irreducible_altdef, intro conjI allI impI) from irr show "?p \ 0" by auto from deg have "degree ?p \ 0" by simp then show "\ ?p dvd 1" unfolding poly_dvd_1 by auto fix b assume "b dvd cf_pos_poly p" also note cf_pos_poly_dvd finally have "b dvd p". with irr[unfolded irreducible_altdef] have "p dvd b \ b dvd 1" by auto then show "?p dvd b \ b dvd 1" by (auto dest: dvd_trans[OF cf_pos_poly_dvd]) qed locale dvd_preserving_hom = comm_semiring_1_hom + assumes hom_eq_mult_hom_imp: "hom x = hom y * hz \ \z. hz = hom z \ x = y * z" begin lemma hom_dvd_hom_iff[simp]: "hom x dvd hom y \ x dvd y" proof assume "hom x dvd hom y" then obtain hz where "hom y = hom x * hz" by (elim dvdE) from hom_eq_mult_hom_imp[OF this] obtain z where "hz = hom z" and mult: "y = x * z" by auto then show "x dvd y" by auto qed auto sublocale unit_preserving_hom proof unfold_locales fix x assume "hom x dvd 1" then have "hom x dvd hom 1" by simp then show "x dvd 1" by (unfold hom_dvd_hom_iff) qed sublocale zero_hom_0 proof (unfold_locales) fix a :: 'a assume "hom a = 0" then have "hom 0 dvd hom a" by auto then have "0 dvd a" by (unfold hom_dvd_hom_iff) then show "a = 0" by auto qed end lemma smult_inverse_monom:"p \ 0 \ smult (inverse c) (p::rat poly) = 1 \ p = [: c :]" proof (cases "c=0") case True thus "p \ 0 \ ?thesis" by auto next case False thus ?thesis by (metis left_inverse right_inverse smult_1 smult_1_left smult_smult) qed lemma of_int_monom:"of_int_poly p = [:rat_of_int c:] \ p = [: c :]" by (induct p, auto) lemma degree_0_content: fixes p :: "int poly" assumes deg: "degree p = 0" shows "content p = abs (coeff p 0)" proof- from deg obtain a where p: "p = [:a:]" by (auto dest: degree0_coeffs) show ?thesis by (auto simp: p) qed lemma prime_elem_imp_gcd_eq: fixes x::"'a:: ring_gcd" shows "prime_elem x \ gcd x y = normalize x \ gcd x y = 1" using prime_elem_imp_coprime [of x y] by (auto simp add: gcd_proj1_iff intro: coprime_imp_gcd_eq_1) lemma irreducible_pos_gcd: fixes p :: "int poly" assumes ir: "irreducible p" and pos: "lead_coeff p > 0" shows "gcd p q \ {1,p}" proof- from pos have "[:sgn (lead_coeff p):] = 1" by auto with prime_elem_imp_gcd_eq[of p, unfolded prime_elem_iff_irreducible, OF ir, of q] show ?thesis by (auto simp: normalize_poly_def) qed lemma irreducible_pos_gcd_twice: fixes p q :: "int poly" assumes p: "irreducible p" "lead_coeff p > 0" and q: "irreducible q" "lead_coeff q > 0" shows "gcd p q = 1 \ p = q" proof (cases "gcd p q = 1") case False note pq = this have "p = gcd p q" using irreducible_pos_gcd [OF p, of q] pq by auto also have "\ = q" using irreducible_pos_gcd [OF q, of p] pq by (auto simp add: ac_simps) finally show ?thesis by auto qed simp interpretation of_rat_hom: field_hom_0' of_rat.. lemma poly_zero_imp_not_unit: assumes "poly p x = 0" shows "\ p dvd 1" proof (rule notI) assume "p dvd 1" from poly_hom.hom_dvd_1[OF this] have "poly p x dvd 1" by auto with assms show False by auto qed lemma poly_prod_mset_zero_iff: fixes x :: "'a :: idom" shows "poly (prod_mset F) x = 0 \ (\f \# F. poly f x = 0)" by (induct F, auto simp: poly_mult_zero_iff) lemma algebraic_imp_represents_irreducible: fixes x :: "'a :: field_char_0" assumes "algebraic x" shows "\p. p represents x \ irreducible p" proof - from assms obtain p where px0: "ipoly p x = 0" and p0: "p \ 0" unfolding algebraic_altdef_ipoly by auto from poly_zero_imp_not_unit[OF px0] have "\ p dvd 1" by (auto dest: of_int_poly_hom.hom_dvd_1[where 'a = 'a]) from mset_factors_exist[OF p0 this] obtain F where F: "mset_factors F p" by auto then have "p = prod_mset F" by auto also have "(of_int_poly ... :: 'a poly) = prod_mset (image_mset of_int_poly F)" by simp finally have "poly ... x = 0" using px0 by auto from this[unfolded poly_prod_mset_zero_iff] obtain f where "f \# F" and fx0: "ipoly f x = 0" by auto with F have "irreducible f" by auto with fx0 show ?thesis by auto qed lemma algebraic_imp_represents_irreducible_cf_pos: assumes "algebraic (x::'a::field_char_0)" shows "\p. p represents x \ irreducible p \ lead_coeff p > 0 \ primitive p" proof - from algebraic_imp_represents_irreducible[OF assms(1)] obtain p where px: "p represents x" and irr: "irreducible p" by auto let ?p = "cf_pos_poly p" from px irr represents_imp_degree have 1: "?p represents x" and 2: "irreducible ?p" and 3: "cf_pos ?p" by (auto intro: irreducible_cf_pos_poly) then show ?thesis by (auto intro: exI[of _ ?p] simp: cf_pos_def) qed lemma gcd_of_int_poly: "gcd (of_int_poly f) (of_int_poly g :: 'a :: {field_char_0,field_gcd} poly) = smult (inverse (of_int (lead_coeff (gcd f g)))) (of_int_poly (gcd f g))" proof - let ?ia = "of_int_poly :: _ \ 'a poly" let ?ir = "of_int_poly :: _ \ rat poly" let ?ra = "map_poly of_rat :: _ \ 'a poly" have id: "?ia x = ?ra (?ir x)" for x by (subst map_poly_map_poly, auto) show ?thesis unfolding id unfolding of_rat_hom.map_poly_gcd[symmetric] unfolding gcd_rat_to_gcd_int by (auto simp: hom_distribs) qed lemma algebraic_imp_represents_unique: fixes x :: "'a :: {field_char_0,field_gcd}" assumes "algebraic x" shows "\! p. p represents x \ irreducible p \ lead_coeff p > 0" (is "Ex1 ?p") proof - from assms obtain p where p: "?p p" and cfp: "cf_pos p" by (auto simp: cf_pos_def dest: algebraic_imp_represents_irreducible_cf_pos) show ?thesis proof (rule ex1I) show "?p p" by fact fix q assume q: "?p q" then have "q represents x" by auto from represents_imp_degree[OF this] q irreducible_content[of q] have cfq: "cf_pos q" by (auto simp: cf_pos_def) show "q = p" proof (rule ccontr) let ?ia = "map_poly of_int :: int poly \ 'a poly" assume "q \ p" with irreducible_pos_gcd_twice[of p q] p q cfp cfq have gcd: "gcd p q = 1" by auto from p q have rt: "ipoly p x = 0" "ipoly q x = 0" unfolding represents_def by auto define c :: 'a where "c = inverse (of_int (lead_coeff (gcd p q)))" have rt: "poly (?ia p) x = 0" "poly (?ia q) x = 0" using rt by auto hence "[:-x,1:] dvd ?ia p" "[:-x,1:] dvd ?ia q" unfolding poly_eq_0_iff_dvd by auto hence "[:-x,1:] dvd gcd (?ia p) (?ia q)" by (rule gcd_greatest) also have "\ = smult c (?ia (gcd p q))" unfolding gcd_of_int_poly c_def .. also have "?ia (gcd p q) = 1" by (simp add: gcd) also have "smult c 1 = [: c :]" by simp finally show False using c_def gcd by (simp add: dvd_iff_poly_eq_0) qed qed qed lemma ipoly_poly_compose: fixes x :: "'a :: idom" shows "ipoly (p \\<^sub>p q) x = ipoly p (ipoly q x)" proof (induct p) case (pCons a p) have "ipoly ((pCons a p) \\<^sub>p q) x = of_int a + ipoly (q * p \\<^sub>p q) x" by (simp add: hom_distribs) also have "ipoly (q * p \\<^sub>p q) x = ipoly q x * ipoly (p \\<^sub>p q) x" by (simp add: hom_distribs) also have "ipoly (p \\<^sub>p q) x = ipoly p (ipoly q x)" unfolding pCons(2) .. also have "of_int a + ipoly q x * \ = ipoly (pCons a p) (ipoly q x)" unfolding map_poly_pCons[OF pCons(1)] by simp finally show ?case . qed simp +lemma algebraic_0[simp]: "algebraic 0" + unfolding algebraic_altdef_ipoly + by (intro exI[of _ "[:0,1:]"], auto) + +lemma algebraic_1[simp]: "algebraic 1" + unfolding algebraic_altdef_ipoly + by (intro exI[of _ "[:-1,1:]"], auto) + + text \Polynomial for unary minus.\ definition poly_uminus :: "'a :: ring_1 poly \ 'a poly" where [code del]: "poly_uminus p \ \i\degree p. monom ((-1)^i * coeff p i) i" lemma poly_uminus_pCons_pCons[simp]: "poly_uminus (pCons a (pCons b p)) = pCons a (pCons (-b) (poly_uminus p))" (is "?l = ?r") proof(cases "p = 0") case False then have deg: "degree (pCons a (pCons b p)) = Suc (Suc (degree p))" by simp show ?thesis by (unfold poly_uminus_def deg sum.atMost_Suc_shift monom_Suc monom_0 sum_pCons_0_commute, simp) next case True then show ?thesis by (auto simp add: poly_uminus_def monom_0 monom_Suc) qed fun poly_uminus_inner :: "'a :: ring_1 list \ 'a poly" where "poly_uminus_inner [] = 0" | "poly_uminus_inner [a] = [:a:]" | "poly_uminus_inner (a#b#cs) = pCons a (pCons (-b) (poly_uminus_inner cs))" lemma poly_uminus_code[code,simp]: "poly_uminus p = poly_uminus_inner (coeffs p)" proof- have "poly_uminus (Poly as) = poly_uminus_inner as" for as :: "'a list" proof (induct "length as" arbitrary:as rule: less_induct) case less show ?case proof(cases as) case Nil then show ?thesis by (simp add: poly_uminus_def) next case [simp]: (Cons a bs) show ?thesis proof (cases bs) case Nil then show ?thesis by (simp add: poly_uminus_def monom_0) next case [simp]: (Cons b cs) show ?thesis by (simp add: less) qed qed qed from this[of "coeffs p"] show ?thesis by simp qed lemma poly_uminus_inner_0[simp]: "poly_uminus_inner as = 0 \ Poly as = 0" by (induct as rule: poly_uminus_inner.induct, auto) lemma degree_poly_uminus_inner[simp]: "degree (poly_uminus_inner as) = degree (Poly as)" by (induct as rule: poly_uminus_inner.induct, auto) lemma ipoly_uminus_inner[simp]: "ipoly (poly_uminus_inner as) (x::'a::comm_ring_1) = ipoly (Poly as) (-x)" by (induct as rule: poly_uminus_inner.induct, auto simp: hom_distribs ring_distribs) lemma represents_uminus: assumes alg: "p represents x" shows "(poly_uminus p) represents (-x)" proof - from representsD[OF alg] have "p \ 0" and rp: "ipoly p x = 0" by auto hence 0: "poly_uminus p \ 0" by simp show ?thesis by (rule representsI[OF _ 0], insert rp, auto) qed lemma content_poly_uminus_inner[simp]: fixes as :: "'a :: ring_gcd list" shows "content (poly_uminus_inner as) = content (Poly as)" by (induct as rule: poly_uminus_inner.induct, auto) text \Multiplicative inverse is represented by @{const reflect_poly}.\ lemma inverse_pow_minus: assumes "x \ (0 :: 'a :: field)" and "i \ n" shows "inverse x ^ n * x ^ i = inverse x ^ (n - i)" using assms by (simp add: field_class.field_divide_inverse power_diff power_inverse) lemma (in inj_idom_hom) reflect_poly_hom: "reflect_poly (map_poly hom p) = map_poly hom (reflect_poly p)" proof - obtain xs where xs: "rev (coeffs p) = xs" by auto show ?thesis unfolding reflect_poly_def coeffs_map_poly_hom rev_map xs by (induct xs, auto simp: hom_distribs) qed lemma ipoly_reflect_poly: assumes x: "(x :: 'a :: field_char_0) \ 0" shows "ipoly (reflect_poly p) x = x ^ (degree p) * ipoly p (inverse x)" (is "?l = ?r") proof - let ?or = "of_int :: int \ 'a" have hom: "inj_idom_hom ?or" .. show ?thesis using poly_reflect_poly_nz[OF x, of "map_poly ?or p"] by (simp add: inj_idom_hom.reflect_poly_hom[OF hom]) qed lemma represents_inverse: assumes x: "x \ 0" and alg: "p represents x" shows "(reflect_poly p) represents (inverse x)" proof (intro representsI) from representsD[OF alg] have "p \ 0" and rp: "ipoly p x = 0" by auto then show "reflect_poly p \ 0" by (metis reflect_poly_0 reflect_poly_at_0_eq_0_iff) show "ipoly (reflect_poly p) (inverse x) = 0" by (subst ipoly_reflect_poly, insert x, auto simp:rp) qed lemma inverse_roots: assumes x: "(x :: 'a :: field_char_0) \ 0" shows "ipoly (reflect_poly p) x = 0 \ ipoly p (inverse x) = 0" using x by (auto simp: ipoly_reflect_poly) context fixes n :: nat begin text \Polynomial for n-th root.\ definition poly_nth_root :: "'a :: idom poly \ 'a poly" where "poly_nth_root p = p \\<^sub>p monom 1 n" lemma ipoly_nth_root: fixes x :: "'a :: idom" shows "ipoly (poly_nth_root p) x = ipoly p (x ^ n)" unfolding poly_nth_root_def ipoly_poly_compose by (simp add: map_poly_monom poly_monom) context assumes n: "n \ 0" begin lemma poly_nth_root_0[simp]: "poly_nth_root p = 0 \ p = 0" unfolding poly_nth_root_def by (rule pcompose_eq_0, insert n, auto simp: degree_monom_eq) lemma represents_nth_root: assumes y: "y^n = x" and alg: "p represents x" shows "(poly_nth_root p) represents y" proof - from representsD[OF alg] have "p \ 0" and rp: "ipoly p x = 0" by auto hence 0: "poly_nth_root p \ 0" by simp show ?thesis by (rule representsI[OF _ 0], unfold ipoly_nth_root y rp, simp) qed lemma represents_nth_root_odd_real: assumes alg: "p represents x" and odd: "odd n" shows "(poly_nth_root p) represents (root n x)" by (rule represents_nth_root[OF odd_real_root_pow[OF odd] alg]) lemma represents_nth_root_pos_real: assumes alg: "p represents x" and pos: "x > 0" shows "(poly_nth_root p) represents (root n x)" proof - from n have id: "Suc (n - 1) = n" by auto show ?thesis proof (rule represents_nth_root[OF _ alg]) show "root n x ^ n = x" using id pos by auto qed qed lemma represents_nth_root_neg_real: assumes alg: "p represents x" and neg: "x < 0" shows "(poly_uminus (poly_nth_root (poly_uminus p))) represents (root n x)" proof - have rt: "root n x = - root n (-x)" unfolding real_root_minus by simp show ?thesis unfolding rt by (rule represents_uminus[OF represents_nth_root_pos_real[OF represents_uminus[OF alg]]], insert neg, auto) qed end end lemma represents_csqrt: assumes alg: "p represents x" shows "(poly_nth_root 2 p) represents (csqrt x)" by (rule represents_nth_root[OF _ _ alg], auto) lemma represents_sqrt: assumes alg: "p represents x" and pos: "x \ 0" shows "(poly_nth_root 2 p) represents (sqrt x)" by (rule represents_nth_root[OF _ _ alg], insert pos, auto) lemma represents_degree: assumes "p represents x" shows "degree p \ 0" proof assume "degree p = 0" from degree0_coeffs[OF this] obtain c where p: "p = [:c:]" by auto from assms[unfolded represents_def p] show False by auto qed text \Polynomial for multiplying a rational number with an algebraic number.\ definition poly_mult_rat_main where "poly_mult_rat_main n d (f :: 'a :: idom poly) = (let fs = coeffs f; k = length fs in poly_of_list (map (\ (fi, i). fi * d ^ i * n ^ (k - Suc i)) (zip fs [0 ..< k])))" definition poly_mult_rat :: "rat \ int poly \ int poly" where "poly_mult_rat r p \ case quotient_of r of (n,d) \ poly_mult_rat_main n d p" lemma coeff_poly_mult_rat_main: "coeff (poly_mult_rat_main n d f) i = coeff f i * n ^ (degree f - i) * d ^ i" proof - have id: "coeff (poly_mult_rat_main n d f) i = (coeff f i * d ^ i) * n ^ (length (coeffs f) - Suc i)" unfolding poly_mult_rat_main_def Let_def poly_of_list_def coeff_Poly unfolding nth_default_coeffs_eq[symmetric] unfolding nth_default_def by auto show ?thesis unfolding id by (simp add: degree_eq_length_coeffs) qed lemma degree_poly_mult_rat_main: "n \ 0 \ degree (poly_mult_rat_main n d f) = (if d = 0 then 0 else degree f)" proof (cases "d = 0") case True thus ?thesis unfolding degree_def unfolding coeff_poly_mult_rat_main by simp next case False hence id: "(d = 0) = False" by simp show "n \ 0 \ ?thesis" unfolding degree_def coeff_poly_mult_rat_main id by (simp add: id) qed lemma ipoly_mult_rat_main: fixes x :: "'a :: {field,ring_char_0}" assumes "d \ 0" and "n \ 0" shows "ipoly (poly_mult_rat_main n d p) x = of_int n ^ degree p * ipoly p (x * of_int d / of_int n)" proof - from assms have d: "(if d = 0 then t else f) = f" for t f :: 'b by simp show ?thesis unfolding poly_altdef of_int_hom.coeff_map_poly_hom mult.assoc[symmetric] of_int_mult[symmetric] sum_distrib_left unfolding of_int_hom.degree_map_poly_hom degree_poly_mult_rat_main[OF assms(2)] d proof (rule sum.cong[OF refl]) fix i assume "i \ {..degree p}" hence i: "i \ degree p" by auto hence id: "of_int n ^ (degree p - i) = (of_int n ^ degree p / of_int n ^ i :: 'a)" by (simp add: assms(2) power_diff) thus "of_int (coeff (poly_mult_rat_main n d p) i) * x ^ i = of_int n ^ degree p * of_int (coeff p i) * (x * of_int d / of_int n) ^ i" unfolding coeff_poly_mult_rat_main by (simp add: field_simps) qed qed lemma degree_poly_mult_rat[simp]: assumes "r \ 0" shows "degree (poly_mult_rat r p) = degree p" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" by auto with assms r have n0: "n \ 0" by simp from quot have id: "poly_mult_rat r p = poly_mult_rat_main n d p" unfolding poly_mult_rat_def by simp show ?thesis unfolding id degree_poly_mult_rat_main[OF n0] using d by simp qed lemma ipoly_mult_rat: assumes r0: "r \ 0" shows "ipoly (poly_mult_rat r p) x = of_int (fst (quotient_of r)) ^ degree p * ipoly p (x * inverse (of_rat r))" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" by auto from r r0 have n: "n \ 0" by simp from r d n have inv: "of_int d / of_int n = inverse r" by simp from quot have id: "poly_mult_rat r p = poly_mult_rat_main n d p" unfolding poly_mult_rat_def by simp show ?thesis unfolding id ipoly_mult_rat_main[OF d n] quot fst_conv of_rat_inverse[symmetric] inv[symmetric] by (simp add: of_rat_divide) qed lemma poly_mult_rat_main_0[simp]: assumes "n \ 0" "d \ 0" shows "poly_mult_rat_main n d p = 0 \ p = 0" proof assume "p = 0" thus "poly_mult_rat_main n d p = 0" by (simp add: poly_mult_rat_main_def) next assume 0: "poly_mult_rat_main n d p = 0" { fix i from 0 have "coeff (poly_mult_rat_main n d p) i = 0" by simp hence "coeff p i = 0" unfolding coeff_poly_mult_rat_main using assms by simp } thus "p = 0" by (intro poly_eqI, auto) qed lemma poly_mult_rat_0[simp]: assumes r0: "r \ 0" shows "poly_mult_rat r p = 0 \ p = 0" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" by auto from r r0 have n: "n \ 0" by simp from quot have id: "poly_mult_rat r p = poly_mult_rat_main n d p" unfolding poly_mult_rat_def by simp show ?thesis unfolding id using n d by simp qed lemma represents_mult_rat: assumes r: "r \ 0" and "p represents x" shows "(poly_mult_rat r p) represents (of_rat r * x)" using assms unfolding represents_def ipoly_mult_rat[OF r] by (simp add: field_simps) text \Polynomial for adding a rational number on an algebraic number. Again, we do not have to factor afterwards.\ definition poly_add_rat :: "rat \ int poly \ int poly" where "poly_add_rat r p \ case quotient_of r of (n,d) \ (poly_mult_rat_main d 1 p \\<^sub>p [:-n,d:])" lemma poly_add_rat_code[code]: "poly_add_rat r p \ case quotient_of r of (n,d) \ let p' = (let fs = coeffs p; k = length fs in poly_of_list (map (\(fi, i). fi * d ^ (k - Suc i)) (zip fs [0..\<^sub>p [:-n,d:] in p''" unfolding poly_add_rat_def poly_mult_rat_main_def Let_def by simp lemma degree_poly_add_rat[simp]: "degree (poly_add_rat r p) = degree p" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" "d > 0" by auto show ?thesis unfolding poly_add_rat_def quot split by (simp add: degree_poly_mult_rat_main d) qed lemma ipoly_add_rat: "ipoly (poly_add_rat r p) x = (of_int (snd (quotient_of r)) ^ degree p) * ipoly p (x - of_rat r)" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" "d > 0" by auto have id: "ipoly [:- n, 1:] (x / of_int d :: 'a) = - of_int n + x / of_int d" by simp show ?thesis unfolding poly_add_rat_def quot split by (simp add: ipoly_mult_rat_main ipoly_poly_compose d r degree_poly_mult_rat_main field_simps id of_rat_divide) qed lemma poly_add_rat_0[simp]: "poly_add_rat r p = 0 \ p = 0" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" "d > 0" by auto show ?thesis unfolding poly_add_rat_def quot split by (simp add: d pcompose_eq_0) qed lemma add_rat_roots: "ipoly (poly_add_rat r p) x = 0 \ ipoly p (x - of_rat r) = 0" unfolding ipoly_add_rat using quotient_of_nonzero by auto lemma represents_add_rat: assumes "p represents x" shows "(poly_add_rat r p) represents (of_rat r + x)" using assms unfolding represents_def ipoly_add_rat by simp (* TODO: move? *) lemmas pos_mult[simplified,simp] = mult_less_cancel_left_pos[of _ 0] mult_less_cancel_left_pos[of _ _ 0] lemma ipoly_add_rat_pos_neg: "ipoly (poly_add_rat r p) (x::'a::linordered_field) < 0 \ ipoly p (x - of_rat r) < 0" "ipoly (poly_add_rat r p) (x::'a::linordered_field) > 0 \ ipoly p (x - of_rat r) > 0" using quotient_of_nonzero unfolding ipoly_add_rat by auto lemma sgn_ipoly_add_rat[simp]: "sgn (ipoly (poly_add_rat r p) (x::'a::linordered_field)) = sgn (ipoly p (x - of_rat r))" (is "sgn ?l = sgn ?r") using ipoly_add_rat_pos_neg[of r p x] by (cases ?r "0::'a" rule: linorder_cases,auto simp: sgn_1_pos sgn_1_neg sgn_eq_0_iff) lemma deg_nonzero_represents: assumes deg: "degree p \ 0" shows "\ x :: complex. p represents x" proof - let ?p = "of_int_poly p :: complex poly" from fundamental_theorem_algebra_factorized[of ?p] obtain as c where id: "smult c (\a\as. [:- a, 1:]) = ?p" and len: "length as = degree ?p" by blast have "degree ?p = degree p" by simp with deg len obtain b bs where as: "as = b # bs" by (cases as, auto) have "p represents b" unfolding represents_def id[symmetric] as using deg by auto thus ?thesis by blast qed end diff --git a/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy @@ -1,1135 +1,991 @@ (* Author: Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Complex Algebraic Numbers\ text \Since currently there is no immediate analog of Sturm's theorem for the complex numbers, we implement complex algebraic numbers via their real and imaginary part. The major algorithm in this theory is a factorization algorithm which factors a rational polynomial over the complex numbers. - This algorithm is then combined with explicit root algorithms to try to factor arbitrary - complex polymials.\ + For factorization of polynomials with complex algebraic coefficients, there + is a separate AFP entry "Factor-Algebraic-Polynomial".\ theory Complex_Algebraic_Numbers imports Real_Roots Complex_Roots_Real_Poly Compare_Complex Jordan_Normal_Form.Char_Poly Berlekamp_Zassenhaus.Code_Abort_Gcd Interval_Arithmetic begin subsection \Complex Roots\ hide_const (open) UnivPoly.coeff hide_const (open) Module.smult hide_const (open) Coset.order abbreviation complex_of_int_poly :: "int poly \ complex poly" where "complex_of_int_poly \ map_poly of_int" abbreviation complex_of_rat_poly :: "rat poly \ complex poly" where "complex_of_rat_poly \ map_poly of_rat" lemma poly_complex_to_real: "(poly (complex_of_int_poly p) (complex_of_real x) = 0) = (poly (real_of_int_poly p) x = 0)" proof - have id: "of_int = complex_of_real o real_of_int" by auto interpret cr: semiring_hom complex_of_real by (unfold_locales, auto) show ?thesis unfolding id by (subst map_poly_map_poly[symmetric], force+) qed lemma represents_cnj: assumes "p represents x" shows "p represents (cnj x)" proof - from assms have p: "p \ 0" and "ipoly p x = 0" by auto hence rt: "poly (complex_of_int_poly p) x = 0" by auto have "poly (complex_of_int_poly p) (cnj x) = 0" by (rule complex_conjugate_root[OF _ rt], subst coeffs_map_poly, auto) with p show ?thesis by auto qed definition poly_2i :: "int poly" where "poly_2i \ [: 4, 0, 1:]" lemma represents_2i: "poly_2i represents (2 * \)" unfolding represents_def poly_2i_def by simp definition root_poly_Re :: "int poly \ int poly" where "root_poly_Re p = cf_pos_poly (poly_mult_rat (inverse 2) (poly_add p p))" lemma root_poly_Re_code[code]: "root_poly_Re p = (let fs = coeffs (poly_add p p); k = length fs in cf_pos_poly (poly_of_list (map (\(fi, i). fi * 2 ^ i) (zip fs [0.. int poly list" where "root_poly_Im p = (let fs = factors_of_int_poly (poly_add p (poly_uminus p)) in remdups ((if (\ f \ set fs. coeff f 0 = 0) then [[:0,1:]] else [])) @ [ cf_pos_poly (poly_div f poly_2i) . f \ fs, coeff f 0 \ 0])" lemma represents_root_poly: assumes "ipoly p x = 0" and p: "p \ 0" shows "(root_poly_Re p) represents (Re x)" and "\ q \ set (root_poly_Im p). q represents (Im x)" proof - let ?Rep = "root_poly_Re p" let ?Imp = "root_poly_Im p" from assms have ap: "p represents x" by auto from represents_cnj[OF this] have apc: "p represents (cnj x)" . from represents_mult_rat[OF _ represents_add[OF ap apc], of "inverse 2"] have "?Rep represents (1 / 2 * (x + cnj x))" unfolding root_poly_Re_def Let_def by (auto simp: hom_distribs) also have "1 / 2 * (x + cnj x) = of_real (Re x)" by (simp add: complex_add_cnj) finally have Rep: "?Rep \ 0" and rt: "ipoly ?Rep (complex_of_real (Re x)) = 0" unfolding represents_def by auto from rt[unfolded poly_complex_to_real] have "ipoly ?Rep (Re x) = 0" . with Rep show "?Rep represents (Re x)" by auto let ?q = "poly_add p (poly_uminus p)" from represents_add[OF ap, of "poly_uminus p" "- cnj x"] represents_uminus[OF apc] have apq: "?q represents (x - cnj x)" by auto from factors_int_poly_represents[OF this] obtain pi where pi: "pi \ set (factors_of_int_poly ?q)" and appi: "pi represents (x - cnj x)" and irr_pi: "irreducible pi" by auto have id: "inverse (2 * \) * (x - cnj x) = of_real (Im x)" apply (cases x) by (simp add: complex_split imaginary_unit.ctr legacy_Complex_simps) from represents_2i have 12: "poly_2i represents (2 * \)" by simp have "\ qi \ set ?Imp. qi represents (inverse (2 * \) * (x - cnj x))" proof (cases "x - cnj x = 0") case False have "poly poly_2i 0 \ 0" unfolding poly_2i_def by auto from represents_div[OF appi 12 this] represents_irr_non_0[OF irr_pi appi False, unfolded poly_0_coeff_0] pi show ?thesis unfolding root_poly_Im_def Let_def by (auto intro: bexI[of _ "cf_pos_poly (poly_div pi poly_2i)"]) next case True hence id2: "Im x = 0" by (simp add: complex_eq_iff) from appi[unfolded True represents_def] have "coeff pi 0 = 0" by (cases pi, auto) with pi have mem: "[:0,1:] \ set ?Imp" unfolding root_poly_Im_def Let_def by auto have "[:0,1:] represents (complex_of_real (Im x))" unfolding id2 represents_def by simp with mem show ?thesis unfolding id by auto qed then obtain qi where qi: "qi \ set ?Imp" "qi \ 0" and rt: "ipoly qi (complex_of_real (Im x)) = 0" unfolding id represents_def by auto from qi rt[unfolded poly_complex_to_real] show "\ qi \ set ?Imp. qi represents (Im x)" by auto qed + + +definition complex_poly :: "int poly \ int poly \ int poly list" where + "complex_poly re im = (let i = [:1,0,1:] + in factors_of_int_poly (poly_add re (poly_mult im i)))" + +lemma complex_poly: assumes re: "re represents (Re x)" + and im: "im represents (Im x)" + shows "\ f \ set (complex_poly re im). f represents x" "\ f. f \ set (complex_poly re im) \ poly_cond f" +proof - + let ?p = "poly_add re (poly_mult im [:1, 0, 1:])" + from re have re: "re represents complex_of_real (Re x)" by simp + from im have im: "im represents complex_of_real (Im x)" by simp + have "[:1,0,1:] represents \" by auto + from represents_add[OF re represents_mult[OF im this]] + have "?p represents of_real (Re x) + complex_of_real (Im x) * \" by simp + also have "of_real (Re x) + complex_of_real (Im x) * \ = x" + by (metis complex_eq mult.commute) + finally have p: "?p represents x" by auto + have "factors_of_int_poly ?p = complex_poly re im" + unfolding complex_poly_def Let_def by simp + from factors_of_int_poly(1)[OF this] factors_of_int_poly(2)[OF this, of x] p + show "\ f \ set (complex_poly re im). f represents x" "\ f. f \ set (complex_poly re im) \ poly_cond f" + unfolding represents_def by auto +qed + +lemma algebraic_complex_iff: "algebraic x = (algebraic (Re x) \ algebraic (Im x))" +proof + assume "algebraic x" + from this[unfolded algebraic_altdef_ipoly] obtain p where "ipoly p x = 0" "p \ 0" by auto + from represents_root_poly[OF this] show "algebraic (Re x) \ algebraic (Im x)" + unfolding represents_def algebraic_altdef_ipoly by auto +next + assume "algebraic (Re x) \ algebraic (Im x)" + from this[unfolded algebraic_altdef_ipoly] obtain re im where + "re represents (Re x)" "im represents (Im x)" by blast + from complex_poly[OF this] show "algebraic x" + unfolding represents_def algebraic_altdef_ipoly by auto +qed + +definition algebraic_complex :: "complex \ bool" where + [simp]: "algebraic_complex = algebraic" + +lemma algebraic_complex_code_unfold[code_unfold]: "algebraic = algebraic_complex" by simp + +lemma algebraic_complex_code[code]: + "algebraic_complex x = (algebraic (Re x) \ algebraic (Im x))" + unfolding algebraic_complex_def algebraic_complex_iff .. + text \Determine complex roots of a polynomial, intended for polynomials of degree 3 or higher, for lower degree polynomials use @{const roots1} or @{const croots2}\ hide_const (open) eq primrec remdups_gen :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "remdups_gen eq [] = []" | "remdups_gen eq (x # xs) = (if (\ y \ set xs. eq x y) then remdups_gen eq xs else x # remdups_gen eq xs)" lemma real_of_3_remdups_equal_3[simp]: "real_of_3 ` set (remdups_gen equal_3 xs) = real_of_3 ` set xs" by (induct xs, auto simp: equal_3) lemma distinct_remdups_equal_3: "distinct (map real_of_3 (remdups_gen equal_3 xs))" by (induct xs, auto, auto simp: equal_3) lemma real_of_3_code [code]: "real_of_3 x = real_of (Real_Alg_Quotient x)" by (transfer, auto) definition "real_parts_3 p = roots_of_3 (root_poly_Re p)" definition "pos_imaginary_parts_3 p = remdups_gen equal_3 (filter (\ x. sgn_3 x = 1) (concat (map roots_of_3 (root_poly_Im p))))" lemma real_parts_3: assumes p: "p \ 0" and "ipoly p x = 0" shows "Re x \ real_of_3 ` set (real_parts_3 p)" unfolding real_parts_3_def using represents_root_poly(1)[OF assms(2,1)] roots_of_3(1) unfolding represents_def by auto lemma distinct_real_parts_3: "distinct (map real_of_3 (real_parts_3 p))" unfolding real_parts_3_def using roots_of_3(2) . lemma pos_imaginary_parts_3: assumes p: "p \ 0" and "ipoly p x = 0" and "Im x > 0" shows "Im x \ real_of_3 ` set (pos_imaginary_parts_3 p)" proof - from represents_root_poly(2)[OF assms(2,1)] obtain q where q: "q \ set (root_poly_Im p)" "q represents Im x" by auto from roots_of_3(1)[of q] have "Im x \ real_of_3 ` set (roots_of_3 q)" using q unfolding represents_def by auto then obtain i3 where i3: "i3 \ set (roots_of_3 q)" and id: "Im x = real_of_3 i3" by auto from \Im x > 0\ have "sgn (Im x) = 1" by simp hence sgn: "sgn_3 i3 = 1" unfolding id by (metis of_rat_eq_1_iff sgn_3) show ?thesis unfolding pos_imaginary_parts_3_def real_of_3_remdups_equal_3 id using sgn i3 q(1) by auto qed lemma distinct_pos_imaginary_parts_3: "distinct (map real_of_3 (pos_imaginary_parts_3 p))" unfolding pos_imaginary_parts_3_def by (rule distinct_remdups_equal_3) lemma remdups_gen_subset: "set (remdups_gen eq xs) \ set xs" by (induct xs, auto) lemma positive_pos_imaginary_parts_3: assumes "x \ set (pos_imaginary_parts_3 p)" shows "0 < real_of_3 x" proof - from subsetD[OF remdups_gen_subset assms[unfolded pos_imaginary_parts_3_def]] have "sgn_3 x = 1" by auto thus ?thesis using sgn_3[of x] by (simp add: sgn_1_pos) qed definition "pair_to_complex ri \ case ri of (r,i) \ Complex (real_of_3 r) (real_of_3 i)" fun get_itvl_2 :: "real_alg_2 \ real interval" where "get_itvl_2 (Irrational n (p,l,r)) = Interval (of_rat l) (of_rat r)" | "get_itvl_2 (Rational r) = (let rr = of_rat r in Interval rr rr)" lemma get_bounds_2: assumes "invariant_2 x" shows "real_of_2 x \\<^sub>i get_itvl_2 x" proof (cases x) case (Irrational n plr) with assms obtain p l r where plr: "plr = (p,l,r)" by (cases plr, auto) from assms Irrational plr have inv1: "invariant_1 (p,l,r)" and id: "real_of_2 x = real_of_1 (p,l,r)" by auto show ?thesis unfolding id using invariant_1D(1)[OF inv1] by (auto simp: plr Irrational) qed (insert assms, auto simp: Let_def) lift_definition get_itvl_3 :: "real_alg_3 \ real interval" is get_itvl_2 . lemma get_itvl_3: "real_of_3 x \\<^sub>i get_itvl_3 x" by (transfer, insert get_bounds_2, auto) fun tighten_bounds_2 :: "real_alg_2 \ real_alg_2" where "tighten_bounds_2 (Irrational n (p,l,r)) = (case tighten_poly_bounds p l r (sgn (ipoly p r)) of (l',r',_) \ Irrational n (p,l',r'))" | "tighten_bounds_2 (Rational r) = Rational r" lemma tighten_bounds_2: assumes inv: "invariant_2 x" shows "real_of_2 (tighten_bounds_2 x) = real_of_2 x" "invariant_2 (tighten_bounds_2 x)" "get_itvl_2 x = Interval l r \ get_itvl_2 (tighten_bounds_2 x) = Interval l' r' \ r' - l' = (r-l) / 2" proof (atomize(full), cases x) case (Irrational n plr) show "real_of_2 (tighten_bounds_2 x) = real_of_2 x \ invariant_2 (tighten_bounds_2 x) \ (get_itvl_2 x = Interval l r \ get_itvl_2 (tighten_bounds_2 x) = Interval l' r' \ r' - l' = (r - l) / 2)" proof - obtain p l r where plr: "plr = (p,l,r)" by (cases plr, auto) let ?tb = "tighten_poly_bounds p l r (sgn (ipoly p r))" obtain l' r' sr' where tb: "?tb = (l',r',sr')" by (cases ?tb, auto) have id: "tighten_bounds_2 x = Irrational n (p,l',r')" unfolding Irrational plr using tb by auto from inv[unfolded Irrational plr] have inv: "invariant_1_2 (p, l, r)" "n = card {y. y \ real_of_1 (p, l, r) \ ipoly p y = 0}" by auto have rof: "real_of_2 x = real_of_1 (p, l, r)" "real_of_2 (tighten_bounds_2 x) = real_of_1 (p, l', r')" using Irrational plr id by auto from inv have inv1: "invariant_1 (p, l, r)" and "poly_cond2 p" by auto hence rc: "\!x. root_cond (p, l, r) x" "poly_cond2 p" by auto note tb' = tighten_poly_bounds[OF tb rc refl] have eq: "real_of_1 (p, l, r) = real_of_1 (p, l', r')" using tb' inv1 using invariant_1_sub_interval(2) by presburger from inv1 tb' have "invariant_1 (p, l', r')" by (metis invariant_1_sub_interval(1)) hence inv2: "invariant_2 (tighten_bounds_2 x)" unfolding id using inv eq by auto thus ?thesis unfolding rof eq unfolding id unfolding Irrational plr using tb'(1-4) arg_cong[OF tb'(5), of real_of_rat] by (auto simp: hom_distribs) qed qed (auto simp: Let_def) lift_definition tighten_bounds_3 :: "real_alg_3 \ real_alg_3" is tighten_bounds_2 using tighten_bounds_2 by auto lemma tighten_bounds_3: "real_of_3 (tighten_bounds_3 x) = real_of_3 x" "get_itvl_3 x = Interval l r \ get_itvl_3 (tighten_bounds_3 x) = Interval l' r' \ r' - l' = (r-l) / 2" by (transfer, insert tighten_bounds_2, auto)+ partial_function (tailrec) filter_list_length :: "('a \ 'a) \ ('a \ bool) \ nat \ 'a list \ 'a list" where [code]: "filter_list_length f p n xs = (let ys = filter p xs in if length ys = n then ys else filter_list_length f p n (map f ys))" lemma filter_list_length: assumes "length (filter P xs) = n" and "\ i x. x \ set xs \ P x \ p ((f ^^ i) x)" and "\ x. x \ set xs \ \ P x \ \ i. \ p ((f ^^ i) x)" and g: "\ x. g (f x) = g x" and P: "\ x. P (f x) = P x" shows "map g (filter_list_length f p n xs) = map g (filter P xs)" proof - from assms(3) have "\ x. \ i. x \ set xs \ \ P x \ \ p ((f ^^ i) x)" by auto from choice[OF this] obtain i where i: "\ x. x \ set xs \ \ P x \ \ p ((f ^^ (i x)) x)" by auto define m where "m = max_list (map i xs)" have m: "\ x. x \ set xs \ \ P x \ \ i \ m. \ p ((f ^^ i) x)" using max_list[of _ "map i xs", folded m_def] i by auto show ?thesis using assms(1-2) m proof (induct m arbitrary: xs rule: less_induct) case (less m xs) define ys where "ys = filter p xs" have xs_ys: "filter P xs = filter P ys" unfolding ys_def filter_filter by (rule filter_cong[OF refl], insert less(3)[of _ 0], auto) have "filter (P \ f) ys = filter P ys" using P unfolding o_def by auto hence id3: "filter P (map f ys) = map f (filter P ys)" unfolding filter_map by simp hence id2: "map g (filter P (map f ys)) = map g (filter P ys)" by (simp add: g) show ?case proof (cases "length ys = n") case True hence id: "filter_list_length f p n xs = ys" unfolding ys_def filter_list_length.simps[of _ _ _ xs] Let_def by auto show ?thesis using True unfolding id xs_ys using less(2) by (metis filter_id_conv length_filter_less less_le xs_ys) next case False { assume "m = 0" from less(4)[unfolded this] have Pp: "x \ set xs \ \ P x \ \ p x" for x by auto with xs_ys False[folded less(2)] have False by (metis (mono_tags, lifting) filter_True mem_Collect_eq set_filter ys_def) } note m0 = this then obtain M where mM: "m = Suc M" by (cases m, auto) hence m: "M < m" by simp from False have id: "filter_list_length f p n xs = filter_list_length f p n (map f ys)" unfolding ys_def filter_list_length.simps[of _ _ _ xs] Let_def by auto show ?thesis unfolding id xs_ys id2[symmetric] proof (rule less(1)[OF m]) fix y assume "y \ set (map f ys)" then obtain x where x: "x \ set xs" "p x" and y: "y = f x" unfolding ys_def by auto { assume "\ P y" hence "\ P x" unfolding y P . from less(4)[OF x(1) this] obtain i where i: "i \ m" and p: "\ p ((f ^^ i) x)" by auto with x obtain j where ij: "i = Suc j" by (cases i, auto) with i have j: "j \ M" unfolding mM by auto have "\ p ((f ^^ j) y)" using p unfolding ij y funpow_Suc_right by simp thus "\i\ M. \ p ((f ^^ i) y)" using j by auto } { fix i assume "P y" hence "P x" unfolding y P . from less(3)[OF x(1) this, of "Suc i"] show "p ((f ^^ i) y)" unfolding y funpow_Suc_right by simp } next show "length (filter P (map f ys)) = n" unfolding id3 length_map using xs_ys less(2) by auto qed qed qed qed definition complex_roots_of_int_poly3 :: "int poly \ complex list" where "complex_roots_of_int_poly3 p \ let n = degree p; rrts = real_roots_of_int_poly p; nr = length rrts; crts = map (\ r. Complex r 0) rrts in if n = nr then crts else let nr_crts = n - nr in if nr_crts = 2 then let pp = real_of_int_poly p div (prod_list (map (\ x. [:-x,1:]) rrts)); cpp = map_poly (\ r. Complex r 0) pp in crts @ croots2 cpp else let nr_pos_crts = nr_crts div 2; rxs = real_parts_3 p; ixs = pos_imaginary_parts_3 p; rts = [(rx, ix). rx <- rxs, ix <- ixs]; crts' = map pair_to_complex (filter_list_length (map_prod tighten_bounds_3 tighten_bounds_3) (\ (r, i). 0 \\<^sub>c ipoly_complex_interval p (Complex_Interval (get_itvl_3 r) (get_itvl_3 i))) nr_pos_crts rts) in crts @ (concat (map (\ x. [x, cnj x]) crts'))" definition complex_roots_of_int_poly_all :: "int poly \ complex list" where "complex_roots_of_int_poly_all p = (let n = degree p in if n \ 3 then complex_roots_of_int_poly3 p else if n = 1 then [roots1 (map_poly of_int p)] else if n = 2 then croots2 (map_poly of_int p) else [])" lemma in_real_itvl_get_bounds_tighten: "real_of_3 x \\<^sub>i get_itvl_3 ((tighten_bounds_3 ^^ n) x)" proof (induct n arbitrary: x) case 0 thus ?case using get_itvl_3[of x] by simp next case (Suc n x) have id: "(tighten_bounds_3 ^^ (Suc n)) x = (tighten_bounds_3 ^^ n) (tighten_bounds_3 x)" by (metis comp_apply funpow_Suc_right) show ?case unfolding id tighten_bounds_3(1)[of x, symmetric] by (rule Suc) qed lemma sandwitch_real: fixes l r :: "nat \ real" assumes la: "l \ a" and ra: "r \ a" and lm: "\i. l i \ m i" and mr: "\i. m i \ r i" shows "m \ a" proof (rule LIMSEQ_I) fix e :: real assume "0 < e" hence e: "0 < e / 2" by simp from LIMSEQ_D[OF la e] obtain n1 where n1: "\ n. n \ n1 \ norm (l n - a) < e/2" by auto from LIMSEQ_D[OF ra e] obtain n2 where n2: "\ n. n \ n2 \ norm (r n - a) < e/2" by auto show "\no. \n\no. norm (m n - a) < e" proof (rule exI[of _ "max n1 n2"], intro allI impI) fix n assume "max n1 n2 \ n" with n1 n2 have *: "norm (l n - a) < e/2" "norm (r n - a) < e/2" by auto from lm[of n] mr[of n] have "norm (m n - a) \ norm (l n - a) + norm (r n - a)" by simp with * show "norm (m n - a) < e" by auto qed qed lemma real_of_tighten_bounds_many[simp]: "real_of_3 ((tighten_bounds_3 ^^ i) x) = real_of_3 x" apply (induct i) using tighten_bounds_3 by auto definition lower_3 where "lower_3 x i \ interval.lower (get_itvl_3 ((tighten_bounds_3 ^^ i) x))" definition upper_3 where "upper_3 x i \ interval.upper (get_itvl_3 ((tighten_bounds_3 ^^ i) x))" lemma interval_size_3: "upper_3 x i - lower_3 x i = (upper_3 x 0 - lower_3 x 0)/2^i" proof (induct i) case (Suc i) have "upper_3 x (Suc i) - lower_3 x (Suc i) = (upper_3 x i - lower_3 x i) / 2" unfolding upper_3_def lower_3_def using tighten_bounds_3 get_itvl_3 by auto with Suc show ?case by auto qed auto lemma interval_size_3_tendsto_0: "(\i. (upper_3 x i - lower_3 x i)) \ 0" by (subst interval_size_3, auto intro: LIMSEQ_divide_realpow_zero) lemma dist_tendsto_0_imp_tendsto: "(\i. \f i - a\ :: real) \ 0 \ f \ a" using LIM_zero_cancel tendsto_rabs_zero_iff by blast lemma upper_3_tendsto: "upper_3 x \ real_of_3 x" proof(rule dist_tendsto_0_imp_tendsto, rule sandwitch_real) fix i obtain l r where lr: "get_itvl_3 ((tighten_bounds_3 ^^ i) x) = Interval l r" by (metis interval.collapse) with get_itvl_3[of "(tighten_bounds_3 ^^ i) x"] show "\(upper_3 x) i - real_of_3 x\ \ (upper_3 x i - lower_3 x i)" unfolding upper_3_def lower_3_def by auto qed (insert interval_size_3_tendsto_0, auto) lemma lower_3_tendsto: "lower_3 x \ real_of_3 x" proof(rule dist_tendsto_0_imp_tendsto, rule sandwitch_real) fix i obtain l r where lr: "get_itvl_3 ((tighten_bounds_3 ^^ i) x) = Interval l r" by (metis interval.collapse) with get_itvl_3[of "(tighten_bounds_3 ^^ i) x"] show "\lower_3 x i - real_of_3 x\ \ (upper_3 x i - lower_3 x i)" unfolding upper_3_def lower_3_def by auto qed (insert interval_size_3_tendsto_0, auto) lemma tends_to_tight_bounds_3: "(\x. get_itvl_3 ((tighten_bounds_3 ^^ x) y)) \\<^sub>i real_of_3 y" using lower_3_tendsto[of y] upper_3_tendsto[of y] unfolding lower_3_def upper_3_def interval_tendsto_def o_def by auto lemma complex_roots_of_int_poly3: assumes p: "p \ 0" and sf: "square_free p" shows "set (complex_roots_of_int_poly3 p) = {x. ipoly p x = 0}" (is "?l = ?r") "distinct (complex_roots_of_int_poly3 p)" proof - interpret map_poly_inj_idom_hom of_real.. define q where "q = real_of_int_poly p" let ?q = "map_poly complex_of_real q" from p have q0: "q \ 0" unfolding q_def by auto hence q: "?q \ 0" by auto define rr where "rr = real_roots_of_int_poly p" define rrts where "rrts = map (\r. Complex r 0) rr" note d = complex_roots_of_int_poly3_def[of p, unfolded Let_def, folded rr_def, folded rrts_def] have rr: "set rr = {x. ipoly p x = 0}" unfolding rr_def using real_roots_of_int_poly(1)[OF p] . have rrts: "set rrts = {x. poly ?q x = 0 \ x \ \}" unfolding rrts_def set_map rr q_def complex_of_real_def[symmetric] by (auto elim: Reals_cases) have dist: "distinct rr" unfolding rr_def using real_roots_of_int_poly(2) . from dist have dist1: "distinct rrts" unfolding rrts_def distinct_map inj_on_def by auto have lrr: "length rr = card {x. poly (real_of_int_poly p) x = 0}" unfolding rr_def using real_roots_of_int_poly[of p] p distinct_card by fastforce have cr: "length rr = card {x. poly ?q x = 0 \ x \ \}" unfolding lrr q_def[symmetric] proof - have "card {x. poly q x = 0} \ card {x. poly (map_poly complex_of_real q) x = 0 \ x \ \}" (is "?l \ ?r") by (rule card_inj_on_le[of of_real], insert poly_roots_finite[OF q], auto simp: inj_on_def) moreover have "?l \ ?r" by (rule card_inj_on_le[of Re, OF _ _ poly_roots_finite[OF q0]], auto simp: inj_on_def elim!: Reals_cases) ultimately show "?l = ?r" by simp qed have conv: "\ x. ipoly p x = 0 \ poly ?q x = 0" unfolding q_def by (subst map_poly_map_poly, auto simp: o_def) have r: "?r = {x. poly ?q x = 0}" unfolding conv .. have "?l = {x. ipoly p x = 0} \ distinct (complex_roots_of_int_poly3 p)" proof (cases "degree p = length rr") case False note oFalse = this show ?thesis proof (cases "degree p - length rr = 2") case False let ?nr = "(degree p - length rr) div 2" define cpxI where "cpxI = pos_imaginary_parts_3 p" define cpxR where "cpxR = real_parts_3 p" let ?rts = "[(rx,ix). rx <- cpxR, ix <- cpxI]" define cpx where "cpx = map pair_to_complex (filter (\ c. ipoly p (pair_to_complex c) = 0) ?rts)" let ?LL = "cpx @ map cnj cpx" let ?LL' = "concat (map (\ x. [x,cnj x]) cpx)" let ?ll = "rrts @ ?LL" let ?ll' = "rrts @ ?LL'" have cpx: "set cpx \ ?r" unfolding cpx_def by auto have ccpx: "cnj ` set cpx \ ?r" using cpx unfolding r by (auto intro!: complex_conjugate_root[of ?q] simp: Reals_def) have "set ?ll \ ?r" using rrts cpx ccpx unfolding r by auto moreover { fix x :: complex assume rt: "ipoly p x = 0" { fix x assume rt: "ipoly p x = 0" and gt: "Im x > 0" define rx where "rx = Re x" let ?x = "Complex rx (Im x)" have x: "x = ?x" by (cases x, auto simp: rx_def) from rt x have rt': "ipoly p ?x = 0" by auto from real_parts_3[OF p rt, folded rx_def] pos_imaginary_parts_3[OF p rt gt] rt' have "?x \ set cpx" unfolding cpx_def cpxI_def cpxR_def by (force simp: pair_to_complex_def[abs_def]) hence "x \ set cpx" using x by simp } note gt = this have cases: "Im x = 0 \ Im x > 0 \ Im x < 0" by auto from rt have rt': "ipoly p (cnj x) = 0" unfolding conv by (intro complex_conjugate_root[of ?q x], auto simp: Reals_def) { assume "Im x > 0" from gt[OF rt this] have "x \ set ?ll" by auto } moreover { assume "Im x < 0" hence "Im (cnj x) > 0" by simp from gt[OF rt' this] have "cnj (cnj x) \ set ?ll" unfolding set_append set_map by blast hence "x \ set ?ll" by simp } moreover { assume "Im x = 0" hence "x \ \" using complex_is_Real_iff by blast with rt rrts have "x \ set ?ll" unfolding conv by auto } ultimately have "x \ set ?ll" using cases by blast } ultimately have lr: "set ?ll = {x. ipoly p x = 0}" by blast let ?rr = "map real_of_3 cpxR" let ?pi = "map real_of_3 cpxI" have dist2: "distinct ?rr" unfolding cpxR_def by (rule distinct_real_parts_3) have dist3: "distinct ?pi" unfolding cpxI_def by (rule distinct_pos_imaginary_parts_3) have idd: "concat (map (map pair_to_complex) (map (\rx. map (Pair rx) cpxI) cpxR)) = concat (map (\r. map (\ i. Complex (real_of_3 r) (real_of_3 i)) cpxI) cpxR)" unfolding pair_to_complex_def by (auto simp: o_def) have dist4: "distinct cpx" unfolding cpx_def proof (rule distinct_map_filter, unfold map_concat idd, unfold distinct_conv_nth, intro allI impI, goal_cases) case (1 i j) from nth_concat_diff[OF 1, unfolded length_map] dist2[unfolded distinct_conv_nth] dist3[unfolded distinct_conv_nth] show ?case by auto qed have dist5: "distinct (map cnj cpx)" using dist4 unfolding distinct_map by (auto simp: inj_on_def) { fix x :: complex have rrts: "x \ set rrts \ Im x = 0" unfolding rrts_def by auto have cpx: "\ x. x \ set cpx \ Im x > 0" unfolding cpx_def cpxI_def by (auto simp: pair_to_complex_def[abs_def] positive_pos_imaginary_parts_3) have cpx': "x \ cnj ` set cpx \ sgn (Im x) = -1" using cpx by auto have "x \ set rrts \ set cpx \ set rrts \ cnj ` set cpx \ set cpx \ cnj ` set cpx" using rrts cpx[of x] cpx' by auto } note dist6 = this have dist: "distinct ?ll" unfolding distinct_append using dist6 by (auto simp: dist1 dist4 dist5) let ?p = "complex_of_int_poly p" have pp: "?p \ 0" using p by auto from p square_free_of_int_poly[OF sf] square_free_rsquarefree have rsf:"rsquarefree ?p" by auto from dist lr have "length ?ll = card {x. poly ?p x = 0}" by (metis distinct_card) also have "\ = degree p" using rsf unfolding rsquarefree_card_degree[OF pp] by simp finally have deg_len: "degree p = length ?ll" by simp let ?P = "\ c. ipoly p (pair_to_complex c) = 0" let ?itvl = "\ r i. ipoly_complex_interval p (Complex_Interval (get_itvl_3 r) (get_itvl_3 i))" let ?itv = "\ (r,i). ?itvl r i" let ?p = "(\ (r,i). 0 \\<^sub>c (?itvl r i))" let ?tb = tighten_bounds_3 let ?f = "map_prod ?tb ?tb" have filter: "map pair_to_complex (filter_list_length ?f ?p ?nr ?rts) = map pair_to_complex (filter ?P ?rts)" proof (rule filter_list_length) have "length (filter ?P ?rts) = length cpx" unfolding cpx_def by simp also have "\ = ?nr" unfolding deg_len by (simp add: rrts_def) finally show "length (filter ?P ?rts) = ?nr" by auto next fix n x assume x: "?P x" obtain r i where xri: "x = (r,i)" by force have id: "(?f ^^ n) x = ((?tb ^^ n) r, (?tb ^^ n) i)" unfolding xri by (induct n, auto) have px: "pair_to_complex x = Complex (real_of_3 r) (real_of_3 i)" unfolding xri pair_to_complex_def by auto show "?p ((?f ^^ n) x)" unfolding id split by (rule ipoly_complex_interval[of "pair_to_complex x" _ p, unfolded x], unfold px, auto simp: in_complex_interval_def in_real_itvl_get_bounds_tighten) next fix x assume x: "x \ set ?rts" "\ ?P x" let ?x = "pair_to_complex x" obtain r i where xri: "x = (r,i)" by force have id: "(?f ^^ n) x = ((?tb ^^ n) r, (?tb ^^ n) i)" for n unfolding xri by (induct n, auto) have px: "?x = Complex (real_of_3 r) (real_of_3 i)" unfolding xri pair_to_complex_def by auto have cvg: "(\ n. ?itv ((?f ^^ n) x)) \\<^sub>c ipoly p ?x" unfolding id split px proof (rule ipoly_complex_interval_tendsto) show "(\ia. Complex_Interval (get_itvl_3 ((?tb ^^ ia) r)) (get_itvl_3 ((?tb ^^ ia) i))) \\<^sub>c Complex (real_of_3 r) (real_of_3 i)" unfolding complex_interval_tendsto_def by (simp add: tends_to_tight_bounds_3 o_def) qed from complex_interval_tendsto_neq[OF this x(2)] show "\ i. \ ?p ((?f ^^ i) x)" unfolding id by auto next show "pair_to_complex (?f x) = pair_to_complex x" for x by (cases x, auto simp: pair_to_complex_def tighten_bounds_3(1)) next show "?P (?f x) = ?P x" for x by (cases x, auto simp: pair_to_complex_def tighten_bounds_3(1)) qed have l: "complex_roots_of_int_poly3 p = ?ll'" unfolding d filter cpx_def[symmetric] cpxI_def[symmetric] cpxR_def[symmetric] using False oFalse by auto have "distinct ?ll' = (distinct rrts \ distinct ?LL' \ set rrts \ set ?LL' = {})" unfolding distinct_append .. also have "set ?LL' = set ?LL" by auto also have "distinct ?LL' = distinct ?LL" by (induct cpx, auto) finally have "distinct ?ll' = distinct ?ll" unfolding distinct_append by auto with dist have "distinct ?ll'" by auto with lr l show ?thesis by auto next case True let ?cr = "map_poly of_real :: real poly \ complex poly" define pp where "pp = complex_of_int_poly p" have id: "pp = map_poly of_real q" unfolding q_def pp_def by (subst map_poly_map_poly, auto simp: o_def) let ?rts = "map (\ x. [:-x,1:]) rr" define rts where "rts = prod_list ?rts" let ?c2 = "?cr (q div rts)" have pq: "\ x. ipoly p x = 0 \ poly q x = 0" unfolding q_def by simp from True have 2: "degree q - card {x. poly q x = 0} = 2" unfolding pq[symmetric] lrr unfolding q_def by simp from True have id: "degree p = length rr \ False" "degree p - length rr = 2 \ True" by auto have l: "?l = of_real ` {x. poly q x = 0} \ set (croots2 ?c2)" unfolding d rts_def id if_False if_True set_append rrts Reals_def by (fold complex_of_real_def q_def, auto) from dist have len_rr: "length rr = card {x. poly q x = 0}" unfolding rr[unfolded pq, symmetric] by (simp add: distinct_card) have rr': "\ r. r \ set rr \ poly q r = 0" using rr unfolding q_def by simp with dist have "q = q div prod_list ?rts * prod_list ?rts" proof (induct rr arbitrary: q) case (Cons r rr q) note dist = Cons(2) let ?p = "q div [:-r,1:]" from Cons.prems(2) have "poly q r = 0" by simp hence "[:-r,1:] dvd q" using poly_eq_0_iff_dvd by blast from dvd_mult_div_cancel[OF this] have "q = ?p * [:-r,1:]" by simp moreover have "?p = ?p div (\x\rr. [:- x, 1:]) * (\x\rr. [:- x, 1:])" proof (rule Cons.hyps) show "distinct rr" using dist by auto fix s assume "s \ set rr" with dist Cons(3) have "s \ r" "poly q s = 0" by auto hence "poly (?p * [:- 1 * r, 1:]) s = 0" using calculation by force thus "poly ?p s = 0" by (simp add: \s \ r\) qed ultimately have q: "q = ?p div (\x\rr. [:- x, 1:]) * (\x\rr. [:- x, 1:]) * [:-r,1:]" by auto also have "\ = (?p div (\x\rr. [:- x, 1:])) * (\x\r # rr. [:- x, 1:])" unfolding mult.assoc by simp also have "?p div (\x\rr. [:- x, 1:]) = q div (\x\r # rr. [:- x, 1:])" unfolding poly_div_mult_right[symmetric] by simp finally show ?case . qed simp hence q_div: "q = q div rts * rts" unfolding rts_def . from q_div q0 have "q div rts \ 0" "rts \ 0" by auto from degree_mult_eq[OF this] have "degree q = degree (q div rts) + degree rts" using q_div by simp also have "degree rts = length rr" unfolding rts_def by (rule degree_linear_factors) also have "\ = card {x. poly q x = 0}" unfolding len_rr by simp finally have deg2: "degree ?c2 = 2" using 2 by simp note croots2 = croots2[OF deg2, symmetric] have "?q = ?cr (q div rts * rts)" using q_div by simp also have "\ = ?cr rts * ?c2" unfolding hom_distribs by simp finally have q_prod: "?q = ?cr rts * ?c2" . from croots2 l have l: "?l = of_real ` {x. poly q x = 0} \ {x. poly ?c2 x = 0}" by simp from r[unfolded q_prod] have r: "?r = {x. poly (?cr rts) x = 0} \ {x. poly ?c2 x = 0}" by auto also have "?cr rts = (\x\rr. ?cr [:- x, 1:])" by (simp add: rts_def o_def of_real_poly_hom.hom_prod_list) also have "{x. poly \ x = 0} = of_real ` set rr" unfolding poly_prod_list_zero_iff by auto also have "set rr = {x. poly q x = 0}" unfolding rr q_def by simp finally have lr: "?l = ?r" unfolding l by simp show ?thesis proof (intro conjI[OF lr]) from sf have sf: "square_free q" unfolding q_def by (rule square_free_of_int_poly) { interpret field_hom_0' complex_of_real .. from sf have "square_free ?q" unfolding square_free_map_poly . } note sf = this have l: "complex_roots_of_int_poly3 p = rrts @ croots2 ?c2" unfolding d rts_def id if_False if_True set_append rrts q_def complex_of_real_def by auto have dist2: "distinct (croots2 ?c2)" unfolding croots2_def Let_def by auto { fix x assume x: "x \ set (croots2 ?c2)" "x \ set rrts" from x(1)[unfolded croots2] have x1: "poly ?c2 x = 0" by auto from x(2) have x2: "poly (?cr rts) x = 0" unfolding rrts_def rts_def complex_of_real_def[symmetric] by (auto simp: poly_prod_list_zero_iff o_def) from square_free_multD(1)[OF sf[unfolded q_prod], of "[: -x, 1:]"] x1 x2 have False unfolding poly_eq_0_iff_dvd by auto } note dist3 = this show "distinct (complex_roots_of_int_poly3 p)" unfolding l distinct_append by (intro conjI dist1 dist2, insert dist3, auto) qed qed next case True have "card {x. poly ?q x = 0} \ degree ?q" by (rule poly_roots_degree[OF q]) also have "\ = degree p" unfolding q_def by simp also have "\ = card {x. poly ?q x = 0 \ x \ \}" using True cr by simp finally have le: "card {x. poly ?q x = 0} \ card {x. poly ?q x = 0 \ x \ \}" by auto have "{x. poly ?q x = 0 \ x \ \} = {x. poly ?q x = 0}" by (rule card_seteq[OF _ _ le], insert poly_roots_finite[OF q], auto) with True rrts dist1 show ?thesis unfolding r d by auto qed thus "distinct (complex_roots_of_int_poly3 p)" "?l = ?r" by auto qed lemma complex_roots_of_int_poly_all: assumes sf: "degree p \ 3 \ square_free p" shows "p \ 0 \ set (complex_roots_of_int_poly_all p) = {x. ipoly p x = 0}" (is "_ \ set ?l = ?r") and "distinct (complex_roots_of_int_poly_all p)" proof - note d = complex_roots_of_int_poly_all_def Let_def have "(p \ 0 \ set ?l = ?r) \ (distinct (complex_roots_of_int_poly_all p))" proof (cases "degree p \ 3") case True hence p: "p \ 0" by auto from True complex_roots_of_int_poly3[OF p] sf show ?thesis unfolding d by auto next case False let ?p = "map_poly (of_int :: int \ complex) p" have deg: "degree ?p = degree p" by (simp add: degree_map_poly) show ?thesis proof (cases "degree p = 1") case True hence l: "?l = [roots1 ?p]" unfolding d by auto from True have "degree ?p = 1" unfolding deg by auto from roots1[OF this] show ?thesis unfolding l roots1_def by auto next case False show ?thesis proof (cases "degree p = 2") case True hence l: "?l = croots2 ?p" unfolding d by auto from True have "degree ?p = 2" unfolding deg by auto from croots2[OF this] show ?thesis unfolding l by (simp add: croots2_def Let_def) next case False with \degree p \ 1\ \degree p \ 2\ \\ (degree p \ 3)\ have True: "degree p = 0" by auto hence l: "?l = []" unfolding d by auto from True have "degree ?p = 0" unfolding deg by auto from roots0[OF _ this] show ?thesis unfolding l by simp qed qed qed thus "p \ 0 \ set ?l = ?r" "distinct (complex_roots_of_int_poly_all p)" by auto qed -text \It now comes the preferred function to compute complex roots of a integer polynomial.\ +text \It now comes the preferred function to compute complex roots of an integer polynomial.\ definition complex_roots_of_int_poly :: "int poly \ complex list" where "complex_roots_of_int_poly p = ( let ps = (if degree p \ 3 then factors_of_int_poly p else [p]) in concat (map complex_roots_of_int_poly_all ps))" definition complex_roots_of_rat_poly :: "rat poly \ complex list" where "complex_roots_of_rat_poly p = complex_roots_of_int_poly (snd (rat_to_int_poly p))" lemma complex_roots_of_int_poly: shows "p \ 0 \ set (complex_roots_of_int_poly p) = {x. ipoly p x = 0}" (is "_ \ ?l = ?r") and "distinct (complex_roots_of_int_poly p)" proof - have "(p \ 0 \ ?l = ?r) \ (distinct (complex_roots_of_int_poly p))" proof (cases "degree p \ 3") case False hence "complex_roots_of_int_poly p = complex_roots_of_int_poly_all p" unfolding complex_roots_of_int_poly_def Let_def by auto with complex_roots_of_int_poly_all[of p] False show ?thesis by auto next case True { fix q assume "q \ set (factors_of_int_poly p)" from factors_of_int_poly(1)[OF refl this] irreducible_imp_square_free[of q] have 0: "q \ 0" and sf: "square_free q" by auto from complex_roots_of_int_poly_all(1)[OF sf 0] complex_roots_of_int_poly_all(2)[OF sf] have "set (complex_roots_of_int_poly_all q) = {x. ipoly q x = 0}" "distinct (complex_roots_of_int_poly_all q)" by auto } note all = this from True have "?l = (\ ((\ p. set (complex_roots_of_int_poly_all p)) ` set (factors_of_int_poly p)))" unfolding complex_roots_of_int_poly_def Let_def by auto also have "\ = (\ ((\ p. {x. ipoly p x = 0}) ` set (factors_of_int_poly p)))" using all by blast finally have l: "?l = (\ ((\ p. {x. ipoly p x = 0}) ` set (factors_of_int_poly p)))" . have lr: "p \ 0 \ ?l = ?r" using l factors_of_int_poly(2)[OF refl, of p] by auto show ?thesis proof (rule conjI[OF lr]) from True have id: "complex_roots_of_int_poly p = concat (map complex_roots_of_int_poly_all (factors_of_int_poly p))" unfolding complex_roots_of_int_poly_def Let_def by auto show "distinct (complex_roots_of_int_poly p)" unfolding id distinct_conv_nth proof (intro allI impI, goal_cases) case (1 i j) let ?fp = "factors_of_int_poly p" let ?rr = "complex_roots_of_int_poly_all" let ?cc = "concat (map ?rr (factors_of_int_poly p))" from nth_concat_diff[OF 1, unfolded length_map] obtain j1 k1 j2 k2 where *: "(j1,k1) \ (j2,k2)" "j1 < length ?fp" "j2 < length ?fp" and "k1 < length (map ?rr ?fp ! j1)" "k2 < length (map ?rr ?fp ! j2)" "?cc ! i = map ?rr ?fp ! j1 ! k1" "?cc ! j = map ?rr ?fp ! j2 ! k2" by blast hence **: "k1 < length (?rr (?fp ! j1))" "k2 < length (?rr (?fp ! j2))" "?cc ! i = ?rr (?fp ! j1) ! k1" "?cc ! j = ?rr (?fp ! j2) ! k2" by auto from * have mem: "?fp ! j1 \ set ?fp" "?fp ! j2 \ set ?fp" by auto show "?cc ! i \ ?cc ! j" proof (cases "j1 = j2") case True with * have "k1 \ k2" by auto with all(2)[OF mem(2)] **(1-2) show ?thesis unfolding **(3-4) unfolding True distinct_conv_nth by auto next case False from \degree p \ 3\ have p: "p \ 0" by auto note fip = factors_of_int_poly(2-3)[OF refl this] show ?thesis unfolding **(3-4) proof define x where "x = ?rr (?fp ! j2) ! k2" assume id: "?rr (?fp ! j1) ! k1 = ?rr (?fp ! j2) ! k2" from ** have x1: "x \ set (?rr (?fp ! j1))" unfolding x_def id[symmetric] by auto from ** have x2: "x \ set (?rr (?fp ! j2))" unfolding x_def by auto from all(1)[OF mem(1)] x1 have x1: "ipoly (?fp ! j1) x = 0" by auto from all(1)[OF mem(2)] x2 have x2: "ipoly (?fp ! j2) x = 0" by auto from False factors_of_int_poly(4)[OF refl, of p] have neq: "?fp ! j1 \ ?fp ! j2" using * unfolding distinct_conv_nth by auto have "poly (complex_of_int_poly p) x = 0" by (meson fip(1) mem(2) x2) from fip(2)[OF this] mem x1 x2 neq show False by blast qed qed qed qed qed thus "p \ 0 \ ?l = ?r" "distinct (complex_roots_of_int_poly p)" by auto qed lemma complex_roots_of_rat_poly: "p \ 0 \ set (complex_roots_of_rat_poly p) = {x. rpoly p x = 0}" (is "_ \ ?l = ?r") "distinct (complex_roots_of_rat_poly p)" proof - obtain c q where cq: "rat_to_int_poly p = (c,q)" by force from rat_to_int_poly[OF this] have pq: "p = smult (inverse (of_int c)) (of_int_poly q)" and c: "c \ 0" by auto show "distinct (complex_roots_of_rat_poly p)" unfolding complex_roots_of_rat_poly_def using complex_roots_of_int_poly(2) . assume p: "p \ 0" with pq c have q: "q \ 0" by auto have id: "{x. rpoly p x = (0 :: complex)} = {x. ipoly q x = 0}" unfolding pq by (simp add: c of_rat_of_int_poly hom_distribs) show "?l = ?r" unfolding complex_roots_of_rat_poly_def cq snd_conv id complex_roots_of_int_poly(1)[OF q] .. qed -definition roots_of_complex_main :: "complex poly \ complex list" where - "roots_of_complex_main p \ let n = degree p in - if n = 0 then [] else if n = 1 then [roots1 p] else if n = 2 then croots2 p - else (complex_roots_of_rat_poly (map_poly to_rat p))" +lemma min_int_poly_complex_of_real[simp]: "min_int_poly (complex_of_real x) = min_int_poly x" +proof (cases "algebraic x") + case False + hence "\ algebraic (complex_of_real x)" unfolding algebraic_complex_iff by auto + with False show ?thesis unfolding min_int_poly_def by auto +next + case True + from min_int_poly_represents[OF True] + have "min_int_poly x represents x" by auto + thus ?thesis + by (intro min_int_poly_unique, auto simp: lead_coeff_min_int_poly_pos) +qed -definition roots_of_complex_poly :: "complex poly \ complex list option" where - "roots_of_complex_poly p \ let (c,pis) = yun_factorization gcd p in - if (c \ 0 \ (\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then - Some (concat (map (roots_of_complex_main o fst) pis)) else None" - -lemma roots_of_complex_main: assumes p: "p \ 0" and deg: "degree p \ 2 \ set (coeffs p) \ \" - shows "set (roots_of_complex_main p) = {x. poly p x = 0}" (is "set ?l = ?r") - and "distinct (roots_of_complex_main p)" -proof - - note d = roots_of_complex_main_def Let_def - have "set ?l = ?r \ distinct (roots_of_complex_main p)" - proof (cases "degree p = 0") - case True - hence "?l = []" unfolding d by auto - with roots0[OF p True] show ?thesis by auto - next - case False note 0 = this - show ?thesis - proof (cases "degree p = 1") - case True - hence "?l = [roots1 p]" unfolding d by auto - with roots1[OF True] show ?thesis by auto - next - case False note 1 = this - show ?thesis - proof (cases "degree p = 2") - case True - hence "?l = croots2 p" unfolding d by auto - with croots2[OF True] show ?thesis by (auto simp: croots2_def Let_def) - next - case False note 2 = this - let ?q = "map_poly to_rat p" - from 0 1 2 have l: "?l = complex_roots_of_rat_poly ?q" unfolding d by auto - from deg 0 1 2 have rat: "set (coeffs p) \ \" by auto - have "p = map_poly (of_rat o to_rat) p" - by (rule sym, rule map_poly_idI, insert rat, auto) - also have "\ = complex_of_rat_poly ?q" - by (subst map_poly_map_poly, auto simp: to_rat) - finally have id: "{x. poly p x = 0} = {x. poly (complex_of_rat_poly ?q) x = 0}" and q: "?q \ 0" - using p by auto - from complex_roots_of_rat_poly[of ?q, folded id l] q - show ?thesis by auto - qed - qed - qed - thus "set ?l = ?r" "distinct ?l" by auto -qed - -lemma roots_of_complex_poly: assumes rt: "roots_of_complex_poly p = Some xs" - shows "set xs = {x. poly p x = 0}" -proof - - obtain c pis where yun: "yun_factorization gcd p = (c,pis)" by force - from rt[unfolded roots_of_complex_poly_def yun split Let_def] - have c: "c \ 0" and pis: "\ p i. (p, i)\ set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" - and xs: "xs = concat (map (roots_of_complex_main \ fst) pis)" - by (auto split: if_splits) - note yun = square_free_factorizationD(1,2,4)[OF yun_factorization(1)[OF yun]] - from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . - have "{x. poly p x = 0} = {x. poly (\(a, i)\set pis. a ^ Suc i) x = 0}" - unfolding p using c by auto - also have "\ = \ ((\ p. {x. poly p x = 0}) ` fst ` set pis)" (is "_ = ?r") - by (subst poly_prod_0, force+) - finally have r: "{x. poly p x = 0} = ?r" . - { - fix p i - assume p: "(p,i) \ set pis" - have "set (roots_of_complex_main p) = {x. poly p x = 0}" - by (rule roots_of_complex_main, insert yun(2)[OF p] pis[OF p], auto) - } note main = this - have "set xs = \ ((\ (p, i). set (roots_of_complex_main p)) ` set pis)" unfolding xs o_def - by auto - also have "\ = ?r" using main by auto - finally show ?thesis unfolding r by simp -qed - -subsection \Factorization of Complex Polynomials\ - -definition factorize_complex_main :: "complex poly \ (complex \ (complex \ nat) list) option" where - "factorize_complex_main p \ let (c,pis) = yun_factorization gcd p in - if ((\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then - Some (c, concat (map (\ (p,i). map (\ r. (r,i)) (roots_of_complex_main p)) pis)) else None" - -definition factorize_complex_poly :: "complex poly \ (complex \ (complex poly \ nat) list) option" where - "factorize_complex_poly p \ map_option - (\ (c,ris). (c, map (\ (r,i). ([:-r,1:],Suc i)) ris)) (factorize_complex_main p)" -lemma factorize_complex_main: assumes rt: "factorize_complex_main p = Some (c,xis)" - shows "p = smult c (\(x, i)\xis. [:- x, 1:] ^ Suc i)" -proof - - obtain d pis where yun: "yun_factorization gcd p = (d,pis)" by force - from rt[unfolded factorize_complex_main_def yun split Let_def] - have pis: "\ p i. (p, i)\set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" - and xis: "xis = concat (map (\(p, i). map (\r. (r, i)) (roots_of_complex_main p)) pis)" - and d: "d = c" - by (auto split: if_splits) - note yun = yun_factorization[OF yun[unfolded d]] - note yun = square_free_factorizationD[OF yun(1)] yun(2)[unfolded snd_conv] - let ?exp = "\ pis. \(x, i)\concat - (map (\(p, i). map (\r. (r, i)) (roots_of_complex_main p)) pis). [:- x, 1:] ^ Suc i" - from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . - also have "(\(a, i)\set pis. a ^ Suc i) = (\(a, i)\pis. a ^ Suc i)" - by (rule prod.distinct_set_conv_list[OF yun(5)]) - also have "\ = ?exp pis" using pis yun(2,6) - proof (induct pis) - case (Cons pi pis) - obtain p i where pi: "pi = (p,i)" by force - let ?rts = "roots_of_complex_main p" - note Cons = Cons[unfolded pi] - have IH: "(\(a, i)\pis. a ^ Suc i) = (?exp pis)" - by (rule Cons(1)[OF Cons(2-4)], auto) - from Cons(2-4)[of p i] have deg: "degree p \ 2 \ (\x\set (coeffs p). x \ \)" - and p: "square_free p" "degree p \ 0" "p \ 0" "monic p" by auto - have "(\(a, i)\(pi # pis). a ^ Suc i) = p ^ Suc i * (\(a, i)\pis. a ^ Suc i)" - unfolding pi by simp - also have "(\(a, i)\pis. a ^ Suc i) = ?exp pis" by (rule IH) - finally have id: "(\(a, i)\(pi # pis). a ^ Suc i) = p ^ Suc i * ?exp pis" by simp - have "?exp (pi # pis) = ?exp [(p,i)] * ?exp pis" unfolding pi by simp - also have "?exp [(p,i)] = (\(x, i)\ (map (\r. (r, i)) ?rts). [:- x, 1:] ^ Suc i)" - by simp - also have "\ = (\ x \ ?rts. [:- x, 1:])^Suc i" - unfolding prod_list_power by (rule arg_cong[of _ _ prod_list], auto) - also have "(\ x \ ?rts. [:- x, 1:]) = p" - proof - - from fundamental_theorem_algebra_factorized[of p, unfolded \monic p\] - obtain as where as: "p = (\a\as. [:- a, 1:])" by auto - also have "\ = (\a\set as. [:- a, 1:])" - proof (rule sym, rule prod.distinct_set_conv_list, rule ccontr) - assume "\ distinct as" - from not_distinct_decomp[OF this] obtain as1 as2 as3 a where - a: "as = as1 @ [a] @ as2 @ [a] @ as3" by blast - define q where "q = (\a\as1 @ as2 @ as3. [:- a, 1:])" - have "p = (\a\as. [:- a, 1:])" by fact - also have "\ = (\a\([a] @ [a]). [:- a, 1:]) * q" - unfolding q_def a map_append prod_list.append by (simp only: ac_simps) - also have "\ = [:-a,1:] * [:-a,1:] * q" by simp - finally have "p = ([:-a,1:] * [:-a,1:]) * q" by simp - hence "[:-a,1:] * [:-a,1:] dvd p" unfolding dvd_def .. - with \square_free p\[unfolded square_free_def, THEN conjunct2, rule_format, of "[:-a,1:]"] - show False by auto - qed - also have "set as = {x. poly p x = 0}" unfolding as poly_prod_list - by (simp add: o_def, induct as, auto) - also have "\ = set ?rts" - by (rule roots_of_complex_main(1)[symmetric], insert p deg, auto) - also have "(\a\set ?rts. [:- a, 1:]) = (\a\?rts. [:- a, 1:])" - by (rule prod.distinct_set_conv_list[OF roots_of_complex_main(2)], insert deg p, auto) - finally show ?thesis by simp +text \TODO: the implemention might be tuned, since the search process should be faster when + using interval arithmetic to figure out the correct factor. + (One might also implement the search via checking @{term "ipoly f x = 0"}, but because of complex-algebraic-number + arithmetic, I think that search would be slower than the current one via "@{term "x \ set (complex_roots_of_int_poly f)"}\ +definition min_int_poly_complex :: "complex \ int poly" where + "min_int_poly_complex x = (if algebraic x then if Im x = 0 then min_int_poly_real (Re x) + else the (find (\ f. x \ set (complex_roots_of_int_poly f)) (complex_poly (min_int_poly (Re x)) (min_int_poly (Im x)))) + else [:0,1:])" + +lemma min_int_poly_complex[code_unfold]: "min_int_poly = min_int_poly_complex" +proof (standard) + fix x + define fs where "fs = complex_poly (min_int_poly (Re x)) (min_int_poly (Im x))" + let ?f = "min_int_poly_complex x" + show "min_int_poly x = ?f" + proof (cases "algebraic x") + case False + thus ?thesis unfolding min_int_poly_def min_int_poly_complex_def by auto + next + case True + show ?thesis + proof (cases "Im x = 0") + case *: True + have id: "?f = min_int_poly_real (Re x)" unfolding min_int_poly_complex_def * using True by auto + show ?thesis unfolding id min_int_poly_real_code_unfold[symmetric] min_int_poly_complex_of_real[symmetric] + using * by (intro arg_cong[of _ _ min_int_poly] complex_eqI, auto) + next + case False + from True[unfolded algebraic_complex_iff] have "algebraic (Re x)" "algebraic (Im x)" by auto + from complex_poly[OF min_int_poly_represents[OF this(1)] min_int_poly_represents[OF this(2)]] + have fs: "\ f \ set fs. ipoly f x = 0" "\ f. f \ set fs \ poly_cond f" unfolding fs_def by auto + let ?fs = "find (\ f. ipoly f x = 0) fs" + let ?fs' = "find (\ f. x \ set (complex_roots_of_int_poly f)) fs" + have "?f = the ?fs'" unfolding min_int_poly_complex_def fs_def + using True False by auto + also have "?fs' = ?fs" + by (rule find_cong[OF refl], subst complex_roots_of_int_poly, insert fs, auto) + finally have id: "?f = the ?fs" . + from fs(1) have "?fs \ None" unfolding find_None_iff by auto + then obtain f where Some: "?fs = Some f" by auto + from find_Some_D[OF this] fs(2)[of f] + show ?thesis unfolding id Some + by (intro min_int_poly_unique, auto) qed - finally have id2: "?exp (pi # pis) = p ^ Suc i * ?exp pis" by simp - show ?case unfolding id id2 .. - qed simp - also have "?exp pis = (\(x, i)\xis. [:- x, 1:] ^ Suc i)" unfolding xis .. - finally show ?thesis unfolding p xis by simp -qed - -lemma distinct_factorize_complex_main: - assumes "factorize_complex_main p = Some fctrs" - shows "distinct (map fst (snd fctrs))" -proof - - from assms have solvable: "\x\set (snd (yun_factorization gcd p)). degree (fst x) \ 2 \ - (\x\set (coeffs (fst x)). x \ \)" - by (auto simp add: factorize_complex_main_def case_prod_unfold - Let_def map_concat o_def split: if_splits) - have sqf: "square_free_factorization p - (fst (yun_factorization gcd p), snd (yun_factorization gcd p))" - by (rule yun_factorization) simp - - have "map fst (snd fctrs) = - concat (map (\x. (roots_of_complex_main (fst x))) (snd (yun_factorization gcd p)))" - using assms by (auto simp add: factorize_complex_main_def case_prod_unfold - Let_def map_concat o_def split: if_splits) - also have "distinct \" - proof (rule distinct_concat, goal_cases) - case 1 - show ?case - proof (subst distinct_map, safe) - from square_free_factorizationD(5)[OF sqf] - show "distinct (snd (yun_factorization gcd p))" . - show "inj_on (\x. (roots_of_complex_main (fst x))) (set (snd (yun_factorization gcd p)))" - proof (rule inj_onI, clarify, goal_cases) - case (1 a1 b1 a2 b2) - { - assume neq: "(a1, b1) \ (a2, b2)" - from 1(1,2)[THEN square_free_factorizationD(2)[OF sqf]] - have "degree a1 \ 0" "degree a2 \ 0" by blast+ - hence [simp]: "a1 \ 0" "a2 \ 0" by auto - from square_free_factorizationD(3)[OF sqf 1(1,2) neq] - have "coprime a1 a2" by simp - from solvable 1(1) have "{z. poly a1 z = 0} = set (roots_of_complex_main a1)" - by (intro roots_of_complex_main(1) [symmetric]) auto - also have "set (roots_of_complex_main a1) = set (roots_of_complex_main a2)" - using 1(3) by (subst (1 2) set_remdups [symmetric]) (simp only: fst_conv) - also from solvable 1(2) have "\ = {z. poly a2 z = 0}" - by (intro roots_of_complex_main) auto - finally have "{z. poly a1 z = 0} = {z. poly a2 z = 0}" . - with coprime_imp_no_common_roots \coprime a1 a2\ - have "{z. poly a1 z = 0} = {}" by auto - with fundamental_theorem_of_algebra constant_degree - have "degree a1 = 0" by auto - with \degree a1 \ 0\ have False by contradiction - } - thus ?case by blast - qed - qed - next - case (2 ys) - then obtain f b where fb: "(f, b) \ set (snd (yun_factorization gcd p))" - and ys: "ys = roots_of_complex_main f" by auto - from square_free_factorizationD(2)[OF sqf fb] have 0: "f \ 0" by auto - from solvable[rule_format, OF fb] have f: "degree f \ 2 \ (set (coeffs f) \ \)" by auto - show ?case unfolding ys - by (rule roots_of_complex_main[OF 0 f]) - next - case (3 ys zs) - then obtain a1 b1 a2 b2 where ab: - "(a1, b1) \ set (snd (yun_factorization gcd p))" - "(a2, b2) \ set (snd (yun_factorization gcd p))" - "ys = roots_of_complex_main a1" "zs = roots_of_complex_main a2" - by auto - with 3 have neq: "(a1,b1) \ (a2,b2)" by auto - from ab(1,2)[THEN square_free_factorizationD(2)[OF sqf]] - have [simp]: "a1 \ 0" "a2 \ 0" by auto - - from square_free_factorizationD(3)[OF sqf ab(1,2) neq] have "coprime a1 a2" by simp - have "set ys = {z. poly a1 z = 0}" "set zs = {z. poly a2 z = 0}" - by (insert solvable ab(1,2), subst ab, - rule roots_of_complex_main; (auto) [])+ - with coprime_imp_no_common_roots \coprime a1 a2\ show ?case by auto - qed - - finally show ?thesis . -qed - -lemma factorize_complex_poly: assumes fp: "factorize_complex_poly p = Some (c,qis)" - shows - "p = smult c (\(q, i)\qis. q ^ i)" - "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" -proof - - from fp[unfolded factorize_complex_poly_def] - obtain pis where fp: "factorize_complex_main p = Some (c,pis)" - and qis: "qis = map (\(r, i). ([:- r, 1:], Suc i)) pis" - by auto - from factorize_complex_main[OF fp] have p: "p = smult c (\(x, i)\pis. [:- x, 1:] ^ Suc i)" . - show "p = smult c (\(q, i)\qis. q ^ i)" unfolding p qis - by (rule arg_cong[of _ _ "\ p. smult c (prod_list p)"], auto) - show "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" - using linear_irreducible_field[of q] unfolding qis by auto + qed qed end diff --git a/thys/Algebraic_Numbers/ROOT b/thys/Algebraic_Numbers/ROOT --- a/thys/Algebraic_Numbers/ROOT +++ b/thys/Algebraic_Numbers/ROOT @@ -1,16 +1,16 @@ chapter AFP session Algebraic_Numbers (AFP) = Berlekamp_Zassenhaus + description "Algebraic Numbers" options [timeout = 3600] sessions Sturm_Sequences theories - Real_Factorization + Complex_Algebraic_Numbers Show_Real_Approx Show_Real_Precise Algebraic_Number_Tests Algebraic_Numbers_External_Code document_files "root.bib" "root.tex" diff --git a/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy @@ -1,3631 +1,3647 @@ (* Author: Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Real Algebraic Numbers\ text \Whereas we previously only proved the closure properties of algebraic numbers, this theory adds the numeric computations that are required to separate the roots, and to pick unique representatives of algebraic numbers. The development is split into three major parts. First, an ambiguous representation of algebraic numbers is used, afterwards another layer is used with special treatment of rational numbers which still does not admit unique representatives, and finally, a quotient type is created modulo the equivalence. The theory also contains a code-setup to implement real numbers via real algebraic numbers.\ text \The results are taken from the textbook \cite[pages 329ff]{AlgNumbers}.\ theory Real_Algebraic_Numbers imports "Abstract-Rewriting.SN_Order_Carrier" Deriving.Compare_Rat Deriving.Compare_Real Jordan_Normal_Form.Gauss_Jordan_IArray_Impl Algebraic_Numbers Sturm_Rat Factors_of_Int_Poly Min_Int_Poly begin text \For algebraic numbers, it turned out that @{const gcd_int_poly} is not preferable to the default implementation of @{const gcd}, which just implements Collin's primitive remainder sequence.\ declare gcd_int_poly_code[code_unfold del] (*TODO: move *) lemma ex1_imp_Collect_singleton: "(\!x. P x) \ P x \ Collect P = {x}" proof(intro iffI conjI, unfold conj_imp_eq_imp_imp) assume "Ex1 P" "P x" then show "Collect P = {x}" by blast next assume Px: "Collect P = {x}" then have "P y \ x = y" for y by auto then show "Ex1 P" by auto from Px show "P x" by auto qed lemma ex1_Collect_singleton[consumes 2]: assumes "\!x. P x" and "P x" and "Collect P = {x} \ thesis" shows thesis by (rule assms(3), subst ex1_imp_Collect_singleton[symmetric], insert assms(1,2), auto) lemma ex1_iff_Collect_singleton: "P x \ (\!x. P x) \ Collect P = {x}" by (subst ex1_imp_Collect_singleton[symmetric], auto) context fixes f assumes bij: "bij f" begin lemma bij_imp_ex1_iff: "(\!x. P (f x)) \ (\!y. P y)" (is "?l = ?r") proof (intro iffI) assume l: ?l then obtain x where "P (f x)" by auto with l have *: "{x} = Collect (P o f)" by auto also have "f ` \ = {y. P (f (Hilbert_Choice.inv f y))}" using bij_image_Collect_eq[OF bij] by auto also have "\ = {y. P y}" proof- have "f (Hilbert_Choice.inv f y) = y" for y by (meson bij bij_inv_eq_iff) then show ?thesis by simp qed finally have "Collect P = {f x}" by auto then show ?r by (fold ex1_imp_Collect_singleton, auto) next assume r: ?r then obtain y where "P y" by auto with r have "{y} = Collect P" by auto also have "Hilbert_Choice.inv f ` \ = Collect (P \ f)" using bij_image_Collect_eq[OF bij_imp_bij_inv[OF bij]] bij by (auto simp: inv_inv_eq) finally have "Collect (P o f) = {Hilbert_Choice.inv f y}" by (simp add: o_def) then show ?l by (fold ex1_imp_Collect_singleton, auto) qed lemma bij_ex1_imp_the_shift: assumes ex1: "\!y. P y" shows "(THE x. P (f x)) = Hilbert_Choice.inv f (THE y. P y)" (is "?l = ?r") proof- from ex1 have "P (THE y. P y)" by (rule the1I2) moreover from ex1[folded bij_imp_ex1_iff] have "P (f (THE x. P (f x)))" by (rule the1I2) ultimately have "(THE y. P y) = f (THE x. P (f x))" using ex1 by auto also have "Hilbert_Choice.inv f \ = (THE x. P (f x))" using bij by (simp add: bij_is_inj) finally show "?l = ?r" by auto qed lemma bij_imp_Collect_image: "{x. P (f x)} = Hilbert_Choice.inv f ` {y. P y}" (is "?l = ?g ` _") proof- have "?l = ?g ` f ` ?l" by (simp add: image_comp inv_o_cancel[OF bij_is_inj[OF bij]]) also have "f ` ?l = {f x | x. P (f x)}" by auto also have "\ = {y. P y}" by (metis bij bij_iff) finally show ?thesis. qed lemma bij_imp_card_image: "card (f ` X) = card X" by (metis bij bij_iff card.infinite finite_imageD inj_onI inj_on_iff_eq_card) end lemma bij_imp_card: assumes bij: "bij f" shows "card {x. P (f x)} = card {x. P x}" unfolding bij_imp_Collect_image[OF bij] bij_imp_card_image[OF bij_imp_bij_inv[OF bij]].. lemma bij_add: "bij (\x. x + y :: 'a :: group_add)" (is ?g1) and bij_minus: "bij (\x. x - y :: 'a)" (is ?g2) and inv_add[simp]: "Hilbert_Choice.inv (\x. x + y) = (\x. x - y)" (is ?g3) and inv_minus[simp]: "Hilbert_Choice.inv (\x. x - y) = (\x. x + y)" (is ?g4) proof- have 1: "(\x. x - y) \ (\x. x + y) = id" and 2: "(\x. x + y) \ (\x. x - y) = id" by auto from o_bij[OF 1 2] show ?g1. from o_bij[OF 2 1] show ?g2. from inv_unique_comp[OF 2 1] show ?g3. from inv_unique_comp[OF 1 2] show ?g4. qed lemmas ex1_shift[simp] = bij_imp_ex1_iff[OF bij_add] bij_imp_ex1_iff[OF bij_minus] lemma ex1_the_shift: assumes ex1: "\!y :: 'a :: group_add. P y" shows "(THE x. P (x + d)) = (THE y. P y) - d" and "(THE x. P (x - d)) = (THE y. P y) + d" unfolding bij_ex1_imp_the_shift[OF bij_add ex1] bij_ex1_imp_the_shift[OF bij_minus ex1] by auto lemma card_shift_image[simp]: shows "card ((\x :: 'a :: group_add. x + d) ` X) = card X" and "card ((\x. x - d) ` X) = card X" by (auto simp: bij_imp_card_image[OF bij_add] bij_imp_card_image[OF bij_minus]) lemma irreducible_root_free: fixes p :: "'a :: {idom,comm_ring_1} poly" assumes irr: "irreducible p" shows "root_free p" proof (cases "degree p" "1::nat" rule: linorder_cases) case greater { fix x assume "poly p x = 0" hence "[:-x,1:] dvd p" using poly_eq_0_iff_dvd by blast then obtain r where p: "p = r * [:-x,1:]" by (elim dvdE, auto) have deg: "degree [:-x,1:] = 1" by simp have dvd: "\ [:-x,1:] dvd 1" by (auto simp: poly_dvd_1) from greater have "degree r \ 0" using degree_mult_le[of r "[:-x,1:]", unfolded deg, folded p] by auto then have "\ r dvd 1" by (auto simp: poly_dvd_1) with p irr irreducibleD[OF irr p] dvd have False by auto } thus ?thesis unfolding root_free_def by auto next case less then have deg: "degree p = 0" by auto from deg obtain p0 where p: "p = [:p0:]" using degree0_coeffs by auto with irr have "p \ 0" by auto with p have "poly p x \ 0" for x by auto thus ?thesis by (auto simp: root_free_def) qed (auto simp: root_free_def) (* **************************************************************** *) subsection \Real Algebraic Numbers -- Innermost Layer\ text \We represent a real algebraic number \\\ by a tuple (p,l,r): \\\ is the unique root in the interval [l,r] and l and r have the same sign. We always assume that p is normalized, i.e., p is the unique irreducible and positive content-free polynomial which represents the algebraic number. This representation clearly admits duplicate representations for the same number, e.g. (...,x-3, 3,3) is equivalent to (...,x-3,2,10).\ subsubsection \Basic Definitions\ type_synonym real_alg_1 = "int poly \ rat \ rat" fun poly_real_alg_1 :: "real_alg_1 \ int poly" where "poly_real_alg_1 (p,_,_) = p" fun rai_ub :: "real_alg_1 \ rat" where "rai_ub (_,_,r) = r" fun rai_lb :: "real_alg_1 \ rat" where "rai_lb (_,l,_) = l" abbreviation "roots_below p x \ {y :: real. y \ x \ ipoly p y = 0}" abbreviation(input) unique_root :: "real_alg_1 \ bool" where "unique_root plr \ (\! x. root_cond plr x)" abbreviation the_unique_root :: "real_alg_1 \ real" where "the_unique_root plr \ (THE x. root_cond plr x)" abbreviation real_of_1 where "real_of_1 \ the_unique_root" lemma root_condI[intro]: assumes "of_rat (rai_lb plr) \ x" and "x \ of_rat (rai_ub plr)" and "ipoly (poly_real_alg_1 plr) x = 0" shows "root_cond plr x" using assms by (auto simp: root_cond_def) lemma root_condE[elim]: assumes "root_cond plr x" and "of_rat (rai_lb plr) \ x \ x \ of_rat (rai_ub plr) \ ipoly (poly_real_alg_1 plr) x = 0 \ thesis" shows thesis using assms by (auto simp: root_cond_def) lemma assumes ur: "unique_root plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" shows unique_rootD: "of_rat l \ x" "x \ of_rat r" "ipoly p x = 0" "root_cond plr x" "x = y \ root_cond plr y" "y = x \ root_cond plr y" and the_unique_root_eqI: "root_cond plr y \ y = x" "root_cond plr y \ x = y" proof - from ur show x: "root_cond plr x" unfolding x_def by (rule theI') have "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) from x[unfolded this] show "of_rat l \ x" "x \ of_rat r" "ipoly p x = 0" by auto from x ur show "root_cond plr y \ y = x" and "root_cond plr y \ x = y" and "x = y \ root_cond plr y" and "y = x \ root_cond plr y" by auto qed lemma unique_rootE: assumes ur: "unique_root plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes main: "of_rat l \ x \ x \ of_rat r \ ipoly p x = 0 \ root_cond plr x \ (\y. x = y \ root_cond plr y) \ (\y. y = x \ root_cond plr y) \ thesis" shows thesis by (rule main, unfold x_def p_def l_def r_def; rule unique_rootD[OF ur]) lemma unique_rootI: assumes "\ y. root_cond plr y \ x = y" "root_cond plr x" shows "unique_root plr" using assms by blast definition poly_cond :: "int poly \ bool" where "poly_cond p = (lead_coeff p > 0 \ irreducible p)" lemma poly_condI[intro]: assumes "lead_coeff p > 0" and "irreducible p" shows "poly_cond p" using assms by (auto simp: poly_cond_def) lemma poly_condD: assumes "poly_cond p" shows "irreducible p" and "lead_coeff p > 0" and "root_free p" and "square_free p" and "p \ 0" using assms unfolding poly_cond_def using irreducible_root_free irreducible_imp_square_free cf_pos_def by auto lemma poly_condE[elim]: assumes "poly_cond p" and "irreducible p \ lead_coeff p > 0 \ root_free p \ square_free p \ p \ 0 \ thesis" shows thesis using assms by (auto dest:poly_condD) definition invariant_1 :: "real_alg_1 \ bool" where "invariant_1 tup \ case tup of (p,l,r) \ unique_root (p,l,r) \ sgn l = sgn r \ poly_cond p" lemma invariant_1I: assumes "unique_root plr" and "sgn (rai_lb plr) = sgn (rai_ub plr)" and "poly_cond (poly_real_alg_1 plr)" shows "invariant_1 plr" using assms by (auto simp: invariant_1_def) lemma assumes "invariant_1 plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" shows invariant_1D: "root_cond plr x" "sgn l = sgn r" "sgn x = of_rat (sgn r)" "unique_root plr" "poly_cond p" "degree p > 0" "primitive p" and invariant_1_root_cond: "\ y. root_cond plr y \ y = x" proof - let ?l = "of_rat l :: real" let ?r = "of_rat r :: real" have plr: "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) from assms show ur: "unique_root plr" and sgn: "sgn l = sgn r" and pc: "poly_cond p" by (auto simp: invariant_1_def) from ur show rc: "root_cond plr x" by (auto simp add: x_def plr intro: theI') from this[unfolded plr] have x: "ipoly p x = 0" and bnd: "?l \ x" "x \ ?r" by auto show "sgn x = of_rat (sgn r)" proof (cases "0::real" "x" rule:linorder_cases) case less with bnd(2) have "0 < ?r" by arith thus ?thesis using less by simp next case equal with bnd have "?l \ 0" "?r \ 0" by auto hence "l \ 0" "r \ 0" by auto with \sgn l = sgn r\ have "l = 0" "r = 0" unfolding sgn_rat_def by (auto split: if_splits) with rc[unfolded plr] show ?thesis by auto next case greater with bnd(1) have "?l < 0" by arith thus ?thesis unfolding \sgn l = sgn r\[symmetric] using greater by simp qed from the_unique_root_eqI[OF ur] rc show "\ y. root_cond plr y \ y = x" by metis { assume "degree p = 0" with poly_zero[OF x, simplified] sgn bnd have "p = 0" by auto with pc have "False" by auto } then show "degree p > 0" by auto with pc show "primitive p" by (intro irreducible_imp_primitive, auto) qed lemma invariant_1E[elim]: assumes "invariant_1 plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes main: "root_cond plr x \ sgn l = sgn r \ sgn x = of_rat (sgn r) \ unique_root plr \ poly_cond p \ degree p > 0 \ primitive p \ thesis" shows thesis apply (rule main) using assms(1) unfolding x_def p_def l_def r_def by (auto dest: invariant_1D) lemma invariant_1_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and "sgn l = sgn r" and ur: "unique_root plr" and "poly_cond p" shows "invariant_1 plr \ real_of_1 plr = x" using the_unique_root_eqI[OF ur x] assms by (cases plr, auto intro: invariant_1I) lemma real_of_1_0: assumes "invariant_1 (p,l,r)" shows [simp]: "the_unique_root (p,l,r) = 0 \ r = 0" and [dest]: "l = 0 \ r = 0" and [intro]: "r = 0 \ l = 0" using assms by (auto simp: sgn_0_0) lemma invariant_1_pos: assumes rc: "invariant_1 (p,l,r)" shows [simp]:"the_unique_root (p,l,r) > 0 \ r > 0" (is "?x > 0 \ _") and [simp]:"the_unique_root (p,l,r) < 0 \ r < 0" and [simp]:"the_unique_root (p,l,r) \ 0 \ r \ 0" and [simp]:"the_unique_root (p,l,r) \ 0 \ r \ 0" and [intro]: "r > 0 \ l > 0" and [dest]: "l > 0 \ r > 0" and [intro]: "r < 0 \ l < 0" and [dest]: "l < 0 \ r < 0" proof(atomize(full),goal_cases) case 1 let ?r = "real_of_rat" from assms[unfolded invariant_1_def] have ur: "unique_root (p,l,r)" and sgn: "sgn l = sgn r" by auto from unique_rootD(1-2)[OF ur] have le: "?r l \ ?x" "?x \ ?r r" by auto from rc show ?case proof (cases r "0::rat" rule:linorder_cases) case greater with sgn have "sgn l = 1" by simp hence l0: "l > 0" by (auto simp: sgn_1_pos) hence "?r l > 0" by auto hence "?x > 0" using le(1) by arith with greater l0 show ?thesis by auto next case equal with real_of_1_0[OF rc] show ?thesis by auto next case less hence "?r r < 0" by auto with le(2) have "?x < 0" by arith with less sgn show ?thesis by (auto simp: sgn_1_neg) qed qed definition invariant_1_2 where "invariant_1_2 rai \ invariant_1 rai \ degree (poly_real_alg_1 rai) > 1" definition poly_cond2 where "poly_cond2 p \ poly_cond p \ degree p > 1" lemma poly_cond2I[intro!]: "poly_cond p \ degree p > 1 \ poly_cond2 p" by (simp add: poly_cond2_def) lemma poly_cond2D: assumes "poly_cond2 p" shows "poly_cond p" and "degree p > 1" using assms by (auto simp: poly_cond2_def) lemma poly_cond2E[elim!]: assumes "poly_cond2 p" and "poly_cond p \ degree p > 1 \ thesis" shows thesis using assms by (auto simp: poly_cond2_def) lemma invariant_1_2_poly_cond2: "invariant_1_2 rai \ poly_cond2 (poly_real_alg_1 rai)" unfolding invariant_1_def invariant_1_2_def poly_cond2_def by auto lemma invariant_1_2I[intro!]: assumes "invariant_1 rai" and "degree (poly_real_alg_1 rai) > 1" shows "invariant_1_2 rai" using assms by (auto simp: invariant_1_2_def) lemma invariant_1_2E[elim!]: assumes "invariant_1_2 rai" and "invariant_1 rai \ degree (poly_real_alg_1 rai) > 1 \ thesis" shows thesis using assms[unfolded invariant_1_2_def] by auto lemma invariant_1_2_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and sgn: "sgn l = sgn r" and ur: "unique_root plr" and p: "poly_cond2 p" shows "invariant_1_2 plr \ real_of_1 plr = x" using invariant_1_realI[OF x] p sgn ur unfolding p_def l_def r_def by auto subsection \Real Algebraic Numbers = Rational + Irrational Real Algebraic Numbers\ text \In the next representation of real algebraic numbers, we distinguish between rational and irrational numbers. The advantage is that whenever we only work on rational numbers, there is not much overhead involved in comparison to the existing implementation of real numbers which just supports the rational numbers. For irrational numbers we additionally store the number of the root, counting from left to right. For instance $-\sqrt{2}$ and $\sqrt{2}$ would be root number 1 and 2 of $x^2 - 2$.\ subsubsection \Definitions and Algorithms on Raw Type\ datatype real_alg_2 = Rational rat | Irrational nat real_alg_1 fun invariant_2 :: "real_alg_2 \ bool" where "invariant_2 (Irrational n rai) = (invariant_1_2 rai \ n = card(roots_below (poly_real_alg_1 rai) (real_of_1 rai)))" | "invariant_2 (Rational r) = True" fun real_of_2 :: "real_alg_2 \ real" where "real_of_2 (Rational r) = of_rat r" | "real_of_2 (Irrational n rai) = real_of_1 rai" definition of_rat_2 :: "rat \ real_alg_2" where [code_unfold]: "of_rat_2 = Rational" lemma of_rat_2: "real_of_2 (of_rat_2 x) = of_rat x" "invariant_2 (of_rat_2 x)" by (auto simp: of_rat_2_def) (* Invariant type *) typedef real_alg_3 = "Collect invariant_2" morphisms rep_real_alg_3 Real_Alg_Invariant by (rule exI[of _ "Rational 0"], auto) setup_lifting type_definition_real_alg_3 lift_definition real_of_3 :: "real_alg_3 \ real" is real_of_2 . (* *************** *) subsubsection \Definitions and Algorithms on Quotient Type\ quotient_type real_alg = real_alg_3 / "\ x y. real_of_3 x = real_of_3 y" morphisms rep_real_alg Real_Alg_Quotient by (auto simp: equivp_def) metis (* real_of *) lift_definition real_of :: "real_alg \ real" is real_of_3 . lemma real_of_inj: "(real_of x = real_of y) = (x = y)" by (transfer, simp) (* ********************** *) subsubsection \Sign\ definition sgn_1 :: "real_alg_1 \ rat" where "sgn_1 x = sgn (rai_ub x)" lemma sgn_1: "invariant_1 x \ real_of_rat (sgn_1 x) = sgn (real_of_1 x)" unfolding sgn_1_def by auto lemma sgn_1_inj: "invariant_1 x \ invariant_1 y \ real_of_1 x = real_of_1 y \ sgn_1 x = sgn_1 y" by (auto simp: sgn_1_def elim!: invariant_1E) (* ********************** *) subsubsection \Normalization: Bounds Close Together\ lemma unique_root_lr: assumes ur: "unique_root plr" shows "rai_lb plr \ rai_ub plr" (is "?l \ ?r") proof - let ?p = "poly_real_alg_1 plr" from ur[unfolded root_cond_def] have ex1: "\! x :: real. of_rat ?l \ x \ x \ of_rat ?r \ ipoly ?p x = 0" by (cases plr, simp) then obtain x :: real where bnd: "of_rat ?l \ x" "x \ of_rat ?r" and rt: "ipoly ?p x = 0" by auto from bnd have "real_of_rat ?l \ of_rat ?r" by linarith thus "?l \ ?r" by (simp add: of_rat_less_eq) qed locale map_poly_zero_hom_0 = base: zero_hom_0 begin sublocale zero_hom_0 "map_poly hom" by (unfold_locales,auto) end interpretation of_int_poly_hom: map_poly_zero_hom_0 "of_int :: int \ 'a :: {ring_1, ring_char_0}" .. lemma ipoly_roots_finite: "p \ 0 \ finite {x :: 'a :: {idom, ring_char_0}. ipoly p x = 0}" by (rule poly_roots_finite, simp) lemma roots_below_the_unique_root: assumes ur: "unique_root (p,l,r)" shows "roots_below p (the_unique_root (p,l,r)) = roots_below p (of_rat r)" (is "roots_below p ?x = _") proof- from ur have rc: "root_cond (p,l,r) ?x" by (auto dest!: unique_rootD) with ur have x: "{x. root_cond (p,l,r) x} = {?x}" by (auto intro: the_unique_root_eqI) from rc have "?x \ {y. ?x \ y \ y \ of_rat r \ ipoly p y = 0}" by auto with rc have l1x: "... = {?x}" by (intro equalityI, fold x(1), force, simp add: x) have rb:"roots_below p (of_rat r) = roots_below p ?x \ {y. ?x < y \ y \ of_rat r \ ipoly p y = 0}" using rc by auto have emp: "\x. the_unique_root (p, l, r) < x \ x \ {ra. ?x \ ra \ ra \ real_of_rat r \ ipoly p ra = 0}" using l1x by auto with rb show ?thesis by auto qed lemma unique_root_sub_interval: assumes ur: "unique_root (p,l,r)" and rc: "root_cond (p,l',r') (the_unique_root (p,l,r))" and between: "l \ l'" "r' \ r" shows "unique_root (p,l',r')" and "the_unique_root (p,l',r') = the_unique_root (p,l,r)" proof - from between have ord: "real_of_rat l \ of_rat l'" "real_of_rat r' \ of_rat r" by (auto simp: of_rat_less_eq) from rc have lr': "real_of_rat l' \ of_rat r'" by auto with ord have lr: "real_of_rat l \ real_of_rat r" by auto show "\!x. root_cond (p, l', r') x" proof (rule, rule rc) fix y assume "root_cond (p,l',r') y" with ord have "root_cond (p,l,r) y" by (auto intro!:root_condI) from the_unique_root_eqI[OF ur this] show "y = the_unique_root (p,l,r)" by simp qed from the_unique_root_eqI[OF this rc] show "the_unique_root (p,l',r') = the_unique_root (p,l,r)" by simp qed lemma invariant_1_sub_interval: assumes rc: "invariant_1 (p,l,r)" and sub: "root_cond (p,l',r') (the_unique_root (p,l,r))" and between: "l \ l'" "r' \ r" shows "invariant_1 (p,l',r')" and "real_of_1 (p,l',r') = real_of_1 (p,l,r)" proof - let ?r = real_of_rat note rcD = invariant_1D[OF rc] from rc have ur: "unique_root (p, l', r')" and id: "the_unique_root (p, l', r') = the_unique_root (p, l, r)" by (atomize(full), intro conjI unique_root_sub_interval[OF _ sub between], auto) show "real_of_1 (p,l',r') = real_of_1 (p,l,r)" using id by simp from rcD(1)[unfolded split] have "?r l \ ?r r" by auto hence lr: "l \ r" by (auto simp: of_rat_less_eq) from unique_rootD[OF ur] have "?r l' \ ?r r'" by auto hence lr': "l' \ r'" by (auto simp: of_rat_less_eq) have "sgn l' = sgn r'" proof (cases "r" "0::rat" rule: linorder_cases) case less with lr lr' between have "l < 0" "l' < 0" "r' < 0" "r < 0" by auto thus ?thesis unfolding sgn_rat_def by auto next case equal with rcD(2) have "l = 0" using sgn_0_0 by auto with equal between lr' have "l' = 0" "r' = 0" by auto then show ?thesis by auto next case greater with rcD(4) have "sgn r = 1" unfolding sgn_rat_def by (cases "r = 0", auto) with rcD(2) have "sgn l = 1" by simp hence l: "l > 0" unfolding sgn_rat_def by (cases "l = 0"; cases "l < 0"; auto) with lr lr' between have "l > 0" "l' > 0" "r' > 0" "r > 0" by auto thus ?thesis unfolding sgn_rat_def by auto qed with between ur rc show "invariant_1 (p,l',r')" by (auto simp add: invariant_1_def id) qed lemma root_sign_change: assumes p0: "poly (p::real poly) x = 0" and pd_ne0: "poly (pderiv p) x \ 0" obtains d where "0 < d" "sgn (poly p (x - d)) \ sgn (poly p (x + d))" "sgn (poly p (x - d)) \ 0" "0 \ sgn (poly p (x + d))" "\ d' > 0. d' \ d \ sgn (poly p (x + d')) = sgn (poly p (x + d)) \ sgn (poly p (x - d')) = sgn (poly p (x - d))" proof - assume a:"(\d. 0 < d \ sgn (poly p (x - d)) \ sgn (poly p (x + d)) \ sgn (poly p (x - d)) \ 0 \ 0 \ sgn (poly p (x + d)) \ \d'>0. d' \ d \ sgn (poly p (x + d')) = sgn (poly p (x + d)) \ sgn (poly p (x - d')) = sgn (poly p (x - d)) \ thesis)" from pd_ne0 consider "poly (pderiv p) x > 0" | "poly (pderiv p) x < 0" by linarith thus ?thesis proof(cases) case 1 obtain d1 where d1:"\h. 0 h < d1 \ poly p (x - h) < 0" "d1 > 0" using DERIV_pos_inc_left[OF poly_DERIV 1] p0 by auto obtain d2 where d2:"\h. 0 h < d2 \ poly p (x + h) > 0" "d2 > 0" using DERIV_pos_inc_right[OF poly_DERIV 1] p0 by auto have g0:"0 < (min d1 d2) / 2" using d1 d2 by auto hence m1:"min d1 d2 / 2 < d1" and m2:"min d1 d2 / 2 < d2" by auto { fix d assume a1:"0 < d" and a2:"d < min d1 d2" have "sgn (poly p (x - d)) = -1" "sgn (poly p (x + d)) = 1" using d1(1)[OF a1] d2(1)[OF a1] a2 by auto } note d=this show ?thesis by(rule a[OF g0];insert d g0 m1 m2, simp) next case 2 obtain d1 where d1:"\h. 0 h < d1 \ poly p (x - h) > 0" "d1 > 0" using DERIV_neg_dec_left[OF poly_DERIV 2] p0 by auto obtain d2 where d2:"\h. 0 h < d2 \ poly p (x + h) < 0" "d2 > 0" using DERIV_neg_dec_right[OF poly_DERIV 2] p0 by auto have g0:"0 < (min d1 d2) / 2" using d1 d2 by auto hence m1:"min d1 d2 / 2 < d1" and m2:"min d1 d2 / 2 < d2" by auto { fix d assume a1:"0 < d" and a2:"d < min d1 d2" have "sgn (poly p (x - d)) = 1" "sgn (poly p (x + d)) = -1" using d1(1)[OF a1] d2(1)[OF a1] a2 by auto } note d=this show ?thesis by(rule a[OF g0];insert d g0 m1 m2, simp) qed qed lemma rational_root_free_degree_iff: assumes rf: "root_free (map_poly rat_of_int p)" and rt: "ipoly p x = 0" shows "(x \ \) = (degree p = 1)" proof assume "x \ \" then obtain y where x: "x = of_rat y" (is "_ = ?x") unfolding Rats_def by blast from rt[unfolded x] have "poly (map_poly rat_of_int p) y = 0" by simp with rf show "degree p = 1" unfolding root_free_def by auto next assume "degree p = 1" from degree1_coeffs[OF this] obtain a b where p: "p = [:a,b:]" and b: "b \ 0" by auto from rt[unfolded p hom_distribs] have "of_int a + x * of_int b = 0" by auto from arg_cong[OF this, of "\ x. (x - of_int a) / of_int b"] have "x = - of_rat (of_int a) / of_rat (of_int b)" using b by auto also have "\ = of_rat (- of_int a / of_int b)" unfolding of_rat_minus of_rat_divide .. finally show "x \ \" by auto qed lemma rational_poly_cond_iff: assumes "poly_cond p" and "ipoly p x = 0" and "degree p > 1" shows "(x \ \) = (degree p = 1)" proof (rule rational_root_free_degree_iff[OF _ assms(2)]) from poly_condD[OF assms(1)] irreducible_connect_rev[of p] assms(3) have p: "irreducible\<^sub>d p" by auto from irreducible\<^sub>d_int_rat[OF this] have "irreducible (map_poly rat_of_int p)" by simp thus "root_free (map_poly rat_of_int p)" by (rule irreducible_root_free) qed lemma poly_cond_degree_gt_1: assumes "poly_cond p" "degree p > 1" "ipoly p x = 0" shows "x \ \" using rational_poly_cond_iff[OF assms(1,3)] assms(2) by simp lemma poly_cond2_no_rat_root: assumes "poly_cond2 p" shows "ipoly p (real_of_rat x) \ 0" using poly_cond_degree_gt_1[of p "real_of_rat x"] assms by auto context fixes p :: "int poly" and x :: "rat" begin lemma gt_rat_sign_change_square_free: assumes ur: "unique_root plr" defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes sf: "square_free p" and in_interval: "l \ y" "y \ r" and py0: "ipoly p y \ 0" and pr0: "ipoly p r \ 0" shows "(sgn (ipoly p y) = sgn (ipoly p r)) = (of_rat y > the_unique_root plr)" (is "?gt = _") proof (rule ccontr) have plr[simp]: "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) assume "?gt \ (real_of_rat y > the_unique_root plr)" note a = this[unfolded plr] from py0 have "p \ 0" unfolding irreducible_def by auto hence p0_real: "real_of_int_poly p \ (0::real poly)" by auto let ?p = "real_of_int_poly p" note urD = unique_rootD[OF ur, simplified] let ?ur = "the_unique_root (p, l, r)" let ?r = real_of_rat from in_interval have in':"?r l \ ?r y" "?r y \ ?r r" unfolding of_rat_less_eq by auto from sf square_free_of_int_poly[of p] square_free_rsquarefree have rsf:"rsquarefree ?p" by auto have ur3:"poly ?p ?ur = 0" using urD(3) by simp from ur have "?ur \ of_rat r" by (auto elim!: unique_rootE) moreover from pr0 have "ipoly p (real_of_rat r) \ 0" by auto with ur3 have "real_of_rat r \ real_of_1 (p,l,r)" by force ultimately have "?ur < ?r r" by auto hence ur2: "0 < ?r r - ?ur" by linarith from rsquarefree_roots rsf ur3 have pd_nonz:"poly (pderiv ?p) ?ur \ 0" by auto obtain d where d':"\d'. d'>0 \ d' \ d \ sgn (poly ?p (?ur + d')) = sgn (poly ?p (?ur + d)) \ sgn (poly ?p (?ur - d')) = sgn (poly ?p (?ur - d))" "sgn (poly ?p (?ur - d)) \ sgn (poly ?p (?ur + d))" "sgn (poly ?p (?ur + d)) \ 0" and d_ge_0:"d > 0" by (metis root_sign_change[OF ur3 pd_nonz]) have sr:"sgn (poly ?p (?ur + d)) = sgn (poly ?p (?r r))" proof (cases "?r r - ?ur \ d") case True show ?thesis using d'(1)[OF ur2 True] by auto next case False hence less:"?ur + d < ?r r" by auto show ?thesis proof(rule no_roots_inbetween_imp_same_sign[OF less,rule_format],goal_cases) case (1 x) from ur 1 d_ge_0 have ran: "real_of_rat l \ x" "x \ real_of_rat r" by (auto elim!: unique_rootE) from 1 d_ge_0 have "the_unique_root (p, l, r) \ x" by auto with ur have "\ root_cond (p,l,r) x" by auto with ran show ?case by auto qed qed consider "?r l < ?ur - d" "?r l < ?ur" | "0 < ?ur - ?r l" "?ur - ?r l \ d" | "?ur = ?r l" using urD by argo hence sl:"sgn (poly ?p (?ur - d)) = sgn (poly ?p (?r l)) \ 0 = sgn (poly ?p (?r l))" proof (cases) case 1 have "sgn (poly ?p (?r l)) = sgn (poly ?p (?ur - d))" proof(rule no_roots_inbetween_imp_same_sign[OF 1(1),rule_format],goal_cases) case (1 x) from ur 1 d_ge_0 have ran: "real_of_rat l \ x" "x \ real_of_rat r" by (auto elim!: unique_rootE) from 1 d_ge_0 have "the_unique_root (p, l, r) \ x" by auto with ur have "\ root_cond (p,l,r) x" by auto with ran show ?case by auto qed thus ?thesis by auto next case 2 show ?thesis using d'(1)[OF 2] by simp qed (insert ur3,simp) have diff_sign: "sgn (ipoly p l) \ sgn (ipoly p r)" using d'(2-) sr sl real_of_rat_sgn by auto have ur':"\x. real_of_rat l \ x \ x \ real_of_rat y \ ipoly p x = 0 \ \ (?r y \ the_unique_root (p,l,r))" proof(standard+,goal_cases) case (1 x) { assume id: "the_unique_root (p,l,r) = ?r y" with unique_rootE[OF ur] ur py0 have False by auto } note neq = this have "root_cond (p, l, r) x" unfolding root_cond_def using 1 a ur by (auto elim!: unique_rootE) with conjunct2[OF 1(1)] 1(2-) the_unique_root_eqI[OF ur] show ?case by (auto intro!: neq) qed hence ur'':"\x. real_of_rat y \ x \ x \ real_of_rat r \ poly (real_of_int_poly p) x \ 0 \ \ (?r y \ the_unique_root (p,l,r))" using urD(2,3) by auto have "(sgn (ipoly p y) = sgn (ipoly p r)) = (?r y > the_unique_root (p,l,r))" proof(cases "sgn (ipoly p r) = sgn (ipoly p y)") case True have sgn:"sgn (poly ?p (real_of_rat l)) \ sgn (poly ?p (real_of_rat y))" using True diff_sign by (simp add: real_of_rat_sgn) have ly:"of_rat l < (of_rat y::real)" using in_interval True diff_sign less_eq_rat_def of_rat_less by auto with no_roots_inbetween_imp_same_sign[OF ly,of ?p] sgn ur' True show ?thesis by force next case False hence ne:"sgn (ipoly p (real_of_rat y)) \ sgn (ipoly p (real_of_rat r))" by (simp add: real_of_rat_sgn) have ry:"of_rat y < (of_rat r::real)" using in_interval False diff_sign less_eq_rat_def of_rat_less by auto obtain x where x:"real_of_rat y \ x" "x \ real_of_rat r" "ipoly p x = 0" using no_roots_inbetween_imp_same_sign[OF ry,of ?p] ne by auto hence lx:"real_of_rat l \ x" using in_interval using False a urD by auto have "?ur = x" using x lx ur by (intro the_unique_root_eqI, auto) then show ?thesis using False x by auto qed thus False using diff_sign(1) a py0 by(cases "ipoly p r = 0";auto simp:sgn_0_0) qed lemma gt_rat_sign_change: assumes ur: "unique_root plr" defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes p: "poly_cond2 p" and in_interval: "l \ y" "y \ r" shows "(sgn (ipoly p y) = sgn (ipoly p r)) = (of_rat y > the_unique_root plr)" (is "?gt = _") proof (rule gt_rat_sign_change_square_free[of plr, folded p_def l_def r_def, OF ur _ in_interval]) note nz = poly_cond2_no_rat_root[OF p] from nz[of y] show "ipoly p y \ 0" by auto from nz[of r] show "ipoly p r \ 0" by auto from p have "irreducible p" by auto thus "square_free p" by (rule irreducible_imp_square_free) qed definition tighten_poly_bounds :: "rat \ rat \ rat \ rat \ rat \ rat" where "tighten_poly_bounds l r sr = (let m = (l + r) / 2; sm = sgn (ipoly p m) in if sm = sr then (l,m,sm) else (m,r,sr))" lemma tighten_poly_bounds: assumes res: "tighten_poly_bounds l r sr = (l',r',sr')" and ur: "unique_root (p,l,r)" and p: "poly_cond2 p" and sr: "sr = sgn (ipoly p r)" shows "root_cond (p,l',r') (the_unique_root (p,l,r))" "l \ l'" "l' \ r'" "r' \ r" "(r' - l') = (r - l) / 2" "sr' = sgn (ipoly p r')" proof - let ?x = "the_unique_root (p,l,r)" let ?x' = "the_unique_root (p,l',r')" let ?m = "(l + r) / 2" note d = tighten_poly_bounds_def Let_def from unique_root_lr[OF ur] have lr: "l \ r" by auto thus "l \ l'" "l' \ r'" "r' \ r" "(r' - l') = (r - l) / 2" "sr' = sgn (ipoly p r')" using res sr unfolding d by (auto split: if_splits) hence "l \ ?m" "?m \ r" by auto note le = gt_rat_sign_change[OF ur,simplified,OF p this] note urD = unique_rootD[OF ur] show "root_cond (p,l',r') ?x" proof (cases "sgn (ipoly p ?m) = sgn (ipoly p r)") case *: False with res sr have id: "l' = ?m" "r' = r" unfolding d by auto from *[unfolded le] urD show ?thesis unfolding id by auto next case *: True with res sr have id: "l' = l" "r' = ?m" unfolding d by auto from *[unfolded le] urD show ?thesis unfolding id by auto qed qed partial_function (tailrec) tighten_poly_bounds_epsilon :: "rat \ rat \ rat \ rat \ rat \ rat" where [code]: "tighten_poly_bounds_epsilon l r sr = (if r - l \ x then (l,r,sr) else (case tighten_poly_bounds l r sr of (l',r',sr') \ tighten_poly_bounds_epsilon l' r' sr'))" partial_function (tailrec) tighten_poly_bounds_for_x :: "rat \ rat \ rat \ rat \ rat \ rat" where [code]: "tighten_poly_bounds_for_x l r sr = (if x < l \ r < x then (l, r, sr) else (case tighten_poly_bounds l r sr of (l',r',sr') \ tighten_poly_bounds_for_x l' r' sr'))" lemma tighten_poly_bounds_epsilon: assumes ur: "unique_root (p,l,r)" defines u: "u \ the_unique_root (p,l,r)" assumes p: "poly_cond2 p" and res: "tighten_poly_bounds_epsilon l r sr = (l',r',sr')" and sr: "sr = sgn (ipoly p r)" and x: "x > 0" shows "l \ l'" "r' \ r" "root_cond (p,l',r') u" "r' - l' \ x" "sr' = sgn (ipoly p r')" proof - let ?u = "the_unique_root (p,l,r)" define delta where "delta = x / 2" have delta: "delta > 0" unfolding delta_def using x by auto let ?dist = "\ (l,r,sr). r - l" let ?rel = "inv_image {(x, y). 0 \ y \ delta_gt delta x y} ?dist" note SN = SN_inv_image[OF delta_gt_SN[OF delta], of ?dist] note simps = res[unfolded tighten_poly_bounds_for_x.simps[of l r]] let ?P = "\ (l,r,sr). unique_root (p,l,r) \ u = the_unique_root (p,l,r) \ tighten_poly_bounds_epsilon l r sr = (l',r',sr') \ sr = sgn (ipoly p r) \ l \ l' \ r' \ r \ r' - l' \ x \ root_cond (p,l',r') u \ sr' = sgn (ipoly p r')" have "?P (l,r,sr)" proof (induct rule: SN_induct[OF SN]) case (1 lr) obtain l r sr where lr: "lr = (l,r,sr)" by (cases lr, auto) show ?case unfolding lr split proof (intro impI) assume ur: "unique_root (p, l, r)" and u: "u = the_unique_root (p, l, r)" and res: "tighten_poly_bounds_epsilon l r sr = (l', r', sr')" and sr: "sr = sgn (ipoly p r)" note tur = unique_rootD[OF ur] note simps = tighten_poly_bounds_epsilon.simps[of l r sr] show "l \ l' \ r' \ r \ r' - l' \ x \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" proof (cases "r - l \ x") case True with res[unfolded simps] ur tur(4) u sr show ?thesis by auto next case False hence x: "r - l > x" by auto let ?tight = "tighten_poly_bounds l r sr" obtain L R SR where tight: "?tight = (L,R,SR)" by (cases ?tight, auto) note tighten = tighten_poly_bounds[OF tight[unfolded sr] ur p] from unique_root_sub_interval[OF ur tighten(1-2,4)] p have ur': "unique_root (p,L,R)" "u = the_unique_root (p,L,R)" unfolding u by auto from res[unfolded simps tight] False sr have "tighten_poly_bounds_epsilon L R SR = (l',r',sr')" by auto note IH = 1[of "(L,R,SR)", unfolded tight split lr, rule_format, OF _ ur' this] have "L \ l' \ r' \ R \ r' - l' \ x \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" by (rule IH, insert tighten False, auto simp: delta_gt_def delta_def) thus ?thesis using tighten by auto qed qed qed from this[unfolded split u, rule_format, OF ur refl res sr] show "l \ l'" "r' \ r" "root_cond (p,l',r') u" "r' - l' \ x" "sr' = sgn (ipoly p r')" using u by auto qed lemma tighten_poly_bounds_for_x: assumes ur: "unique_root (p,l,r)" defines u: "u \ the_unique_root (p,l,r)" assumes p: "poly_cond2 p" and res: "tighten_poly_bounds_for_x l r sr = (l',r',sr')" and sr: "sr = sgn (ipoly p r)" shows "l \ l'" "l' \ r'" "r' \ r" "root_cond (p,l',r') u" "\ (l' \ x \ x \ r')" "sr' = sgn (ipoly p r')" "unique_root (p,l',r')" proof - let ?u = "the_unique_root (p,l,r)" let ?x = "real_of_rat x" define delta where "delta = abs ((u - ?x) / 2)" let ?p = "real_of_int_poly p" note ru = unique_rootD[OF ur] { assume "u = ?x" note u = this[unfolded u] from poly_cond2_no_rat_root[OF p] ur have False by (elim unique_rootE, auto simp: u) } hence delta: "delta > 0" unfolding delta_def by auto let ?dist = "\ (l,r,sr). real_of_rat (r - l)" let ?rel = "inv_image {(x, y). 0 \ y \ delta_gt delta x y} ?dist" note SN = SN_inv_image[OF delta_gt_SN[OF delta], of ?dist] note simps = res[unfolded tighten_poly_bounds_for_x.simps[of l r]] let ?P = "\ (l,r,sr). unique_root (p,l,r) \ u = the_unique_root (p,l,r) \ tighten_poly_bounds_for_x l r sr = (l',r',sr') \ sr = sgn (ipoly p r) \ l \ l' \ r' \ r \ \ (l' \ x \ x \ r') \ root_cond (p,l',r') u \ sr' = sgn (ipoly p r')" have "?P (l,r,sr)" proof (induct rule: SN_induct[OF SN]) case (1 lr) obtain l r sr where lr: "lr = (l,r,sr)" by (cases lr, auto) let ?l = "real_of_rat l" let ?r = "real_of_rat r" show ?case unfolding lr split proof (intro impI) assume ur: "unique_root (p, l, r)" and u: "u = the_unique_root (p, l, r)" and res: "tighten_poly_bounds_for_x l r sr = (l', r', sr')" and sr: "sr = sgn (ipoly p r)" note tur = unique_rootD[OF ur] note simps = tighten_poly_bounds_for_x.simps[of l r] show "l \ l' \ r' \ r \ \ (l' \ x \ x \ r') \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" proof (cases "x < l \ r < x") case True with res[unfolded simps] ur tur(4) u sr show ?thesis by auto next case False hence x: "?l \ ?x" "?x \ ?r" by (auto simp: of_rat_less_eq) let ?tight = "tighten_poly_bounds l r sr" obtain L R SR where tight: "?tight = (L,R,SR)" by (cases ?tight, auto) note tighten = tighten_poly_bounds[OF tight ur p sr] from unique_root_sub_interval[OF ur tighten(1-2,4)] p have ur': "unique_root (p,L,R)" "u = the_unique_root (p,L,R)" unfolding u by auto from res[unfolded simps tight] False have "tighten_poly_bounds_for_x L R SR = (l',r',sr')" by auto note IH = 1[of ?tight, unfolded tight split lr, rule_format, OF _ ur' this] let ?DIFF = "real_of_rat (R - L)" let ?diff = "real_of_rat (r - l)" have diff0: "0 \ ?DIFF" using tighten(3) by (metis cancel_comm_monoid_add_class.diff_cancel diff_right_mono of_rat_less_eq of_rat_hom.hom_zero) have *: "r - l - (r - l) / 2 = (r - l) / 2" by (auto simp: field_simps) have "delta_gt delta ?diff ?DIFF = (abs (u - of_rat x) \ real_of_rat (r - l) * 1)" unfolding delta_gt_def tighten(5) delta_def of_rat_diff[symmetric] * by (simp add: hom_distribs) also have "real_of_rat (r - l) * 1 = ?r - ?l" unfolding of_rat_divide of_rat_mult of_rat_diff by auto also have "abs (u - of_rat x) \ ?r - ?l" using x ur by (elim unique_rootE, auto simp: u) finally have delta: "delta_gt delta ?diff ?DIFF" . have "L \ l' \ r' \ R \ \ (l' \ x \ x \ r') \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" by (rule IH, insert delta diff0 tighten(6), auto) with \l \ L\ \R \ r\ show ?thesis by auto qed qed qed from this[unfolded split u, rule_format, OF ur refl res sr] show *: "l \ l'" "r' \ r" "root_cond (p,l',r') u" "\ (l' \ x \ x \ r')" "sr' = sgn (ipoly p r')" unfolding u by auto from *(3)[unfolded split] have "real_of_rat l' \ of_rat r'" by auto thus "l' \ r'" unfolding of_rat_less_eq . show "unique_root (p,l',r')" using ur *(1-3) p poly_condD(5) u unique_root_sub_interval(1) by blast qed end definition real_alg_precision :: rat where "real_alg_precision \ Rat.Fract 1 2" lemma real_alg_precision: "real_alg_precision > 0" by eval definition normalize_bounds_1_main :: "rat \ real_alg_1 \ real_alg_1" where "normalize_bounds_1_main eps rai = (case rai of (p,l,r) \ let (l',r',sr') = tighten_poly_bounds_epsilon p eps l r (sgn (ipoly p r)); fr = rat_of_int (floor r'); (l'',r'',_) = tighten_poly_bounds_for_x p fr l' r' sr' in (p,l'',r''))" definition normalize_bounds_1 :: "real_alg_1 \ real_alg_1" where "normalize_bounds_1 = (normalize_bounds_1_main real_alg_precision)" context fixes p q and l r :: rat assumes cong: "\ x. real_of_rat l \ x \ x \ of_rat r \ (ipoly p x = (0 :: real)) = (ipoly q x = 0)" begin lemma root_cond_cong: "root_cond (p,l,r) = root_cond (q,l,r)" by (intro ext, insert cong, auto simp: root_cond_def) lemma the_unique_root_cong: "the_unique_root (p,l,r) = the_unique_root (q,l,r)" unfolding root_cond_cong .. lemma unique_root_cong: "unique_root (p,l,r) = unique_root (q,l,r)" unfolding root_cond_cong .. end lemma normalize_bounds_1_main: assumes eps: "eps > 0" and rc: "invariant_1_2 x" defines y: "y \ normalize_bounds_1_main eps x" shows "invariant_1_2 y \ (real_of_1 y = real_of_1 x)" proof - obtain p l r where x: "x = (p,l,r)" by (cases x) auto note rc = rc[unfolded x] obtain l' r' sr' where tb: "tighten_poly_bounds_epsilon p eps l r (sgn (ipoly p r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?fr = "rat_of_int (floor r')" obtain l'' r'' sr'' where tbx: "tighten_poly_bounds_for_x p ?fr l' r' sr' = (l'',r'',sr'')" by (cases rule: prod_cases3, auto) from y[unfolded normalize_bounds_1_main_def x] tb tbx have y: "y = (p, l'', r'')" by (auto simp: Let_def) from rc have "unique_root (p, l, r)" and p2: "poly_cond2 p" by auto from tighten_poly_bounds_epsilon[OF this tb refl eps] have bnd: "l \ l'" "r' \ r" and rc': "root_cond (p, l', r') (the_unique_root (p, l, r))" and eps: "r' - l' \ eps" (* currently not relevant for lemma *) and sr': "sr' = sgn (ipoly p r')" by auto from invariant_1_sub_interval[OF _ rc' bnd] rc have inv': "invariant_1 (p, l', r')" and eq: "real_of_1 (p, l', r') = real_of_1 (p, l, r)" by auto have bnd: "l' \ l''" "r'' \ r'" and rc': "root_cond (p, l'', r'') (the_unique_root (p, l', r'))" by (rule tighten_poly_bounds_for_x[OF _ p2 tbx sr'], fact invariant_1D[OF inv'])+ from invariant_1_sub_interval[OF inv' rc' bnd] p2 eq show ?thesis unfolding y x by auto qed lemma normalize_bounds_1: assumes x: "invariant_1_2 x" shows "invariant_1_2 (normalize_bounds_1 x) \ (real_of_1 (normalize_bounds_1 x) = real_of_1 x)" proof(cases x) case xx:(fields p l r) let ?res = "(p,l,r)" have norm: "normalize_bounds_1 x = (normalize_bounds_1_main real_alg_precision ?res)" unfolding normalize_bounds_1_def by (simp add: xx) from x have x: "invariant_1_2 ?res" "real_of_1 ?res = real_of_1 x" unfolding xx by auto from normalize_bounds_1_main[OF real_alg_precision x(1)] x(2-) show ?thesis unfolding normalize_bounds_1_def xx by auto qed lemma normalize_bound_1_poly: "poly_real_alg_1 (normalize_bounds_1 rai) = poly_real_alg_1 rai" unfolding normalize_bounds_1_def normalize_bounds_1_main_def Let_def by (auto split: prod.splits) definition real_alg_2_main :: "root_info \ real_alg_1 \ real_alg_2" where "real_alg_2_main ri rai \ let p = poly_real_alg_1 rai in (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else (case normalize_bounds_1 rai of (p',l,r) \ Irrational (root_info.number_root ri r) (p',l,r)))" definition real_alg_2 :: "real_alg_1 \ real_alg_2" where "real_alg_2 rai \ let p = poly_real_alg_1 rai in (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else (case normalize_bounds_1 rai of (p',l,r) \ Irrational (root_info.number_root (root_info p) r) (p',l,r)))" lemma degree_1_ipoly: assumes "degree p = Suc 0" shows "ipoly p x = 0 \ (x = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1)))" proof - from roots1[of "map_poly real_of_int p"] assms have "ipoly p x = 0 \ x \ {roots1 (real_of_int_poly p)}" by auto also have "\ = (x = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1)))" unfolding Fract_of_int_quotient roots1_def hom_distribs by auto finally show ?thesis . qed lemma invariant_1_degree_0: assumes inv: "invariant_1 rai" shows "degree (poly_real_alg_1 rai) \ 0" (is "degree ?p \ 0") proof (rule notI) assume deg: "degree ?p = 0" from inv have "ipoly ?p (real_of_1 rai) = 0" by auto with deg have "?p = 0" by (meson less_Suc0 representsI represents_degree) with inv show False by auto qed lemma real_alg_2_main: assumes inv: "invariant_1 rai" defines [simp]: "p \ poly_real_alg_1 rai" assumes ric: "irreducible (poly_real_alg_1 rai) \ root_info_cond ri (poly_real_alg_1 rai)" shows "invariant_2 (real_alg_2_main ri rai)" "real_of_2 (real_alg_2_main ri rai) = real_of_1 rai" proof (atomize(full)) define l r where [simp]: "l \ rai_lb rai" and [simp]: "r \ rai_ub rai" show "invariant_2 (real_alg_2_main ri rai) \ real_of_2 (real_alg_2_main ri rai) = real_of_1 rai" unfolding id using invariant_1D proof (cases "degree p" "Suc 0" rule: linorder_cases) case deg: equal hence id: "real_alg_2_main ri rai = Rational (Rat.Fract (- coeff p 0) (coeff p 1))" unfolding real_alg_2_main_def Let_def by auto note rc = invariant_1D[OF inv] from degree_1_ipoly[OF deg, of "the_unique_root rai"] rc(1) show ?thesis unfolding id by auto next case deg: greater with inv have inv: "invariant_1_2 rai" unfolding p_def by auto define rai' where "rai' = normalize_bounds_1 rai" have rai': "real_of_1 rai = real_of_1 rai'" and inv': "invariant_1_2 rai'" unfolding rai'_def using normalize_bounds_1[OF inv] by auto obtain p' l' r' where "rai' = (p',l',r')" by (cases rai') with arg_cong[OF rai'_def, of poly_real_alg_1, unfolded normalize_bound_1_poly] split have split: "rai' = (p,l',r')" by auto from inv'[unfolded split] have "poly_cond p" by auto from poly_condD[OF this] have irr: "irreducible p" by simp from ric irr have ric: "root_info_cond ri p" by auto have id: "real_alg_2_main ri rai = (Irrational (root_info.number_root ri r') rai')" unfolding real_alg_2_main_def Let_def using deg split rai'_def by (auto simp: rai'_def rai') show ?thesis unfolding id using rai' root_info_condD(2)[OF ric] inv'[unfolded split] apply (elim invariant_1_2E invariant_1E) using inv' by(auto simp: split roots_below_the_unique_root) next case deg: less then have "degree p = 0" by auto from this invariant_1_degree_0[OF inv] have "p = 0" by simp with inv show ?thesis by auto qed qed lemma real_alg_2: assumes "invariant_1 rai" shows "invariant_2 (real_alg_2 rai)" "real_of_2 (real_alg_2 rai) = real_of_1 rai" proof - have deg: "0 < degree (poly_real_alg_1 rai)" using assms by auto have "real_alg_2 rai = real_alg_2_main (root_info (poly_real_alg_1 rai)) rai" unfolding real_alg_2_def real_alg_2_main_def Let_def by auto from real_alg_2_main[OF assms root_info, folded this, simplified] deg show "invariant_2 (real_alg_2 rai)" "real_of_2 (real_alg_2 rai) = real_of_1 rai" by auto qed lemma invariant_2_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and sgn: "sgn l = sgn r" and ur: "unique_root plr" and p: "poly_cond p" shows "invariant_2 (real_alg_2 plr) \ real_of_2 (real_alg_2 plr) = x" using invariant_1_realI[OF x,folded p_def l_def r_def] sgn ur p real_alg_2[of plr] by auto (* ********************* *) subsubsection \Comparisons\ fun compare_rat_1 :: "rat \ real_alg_1 \ order" where "compare_rat_1 x (p,l,r) = (if x < l then Lt else if x > r then Gt else if sgn (ipoly p x) = sgn(ipoly p r) then Gt else Lt)" lemma compare_rat_1: assumes rai: "invariant_1_2 y" shows "compare_rat_1 x y = compare (of_rat x) (real_of_1 y)" proof- define p l r where "p \ poly_real_alg_1 y" "l \ rai_lb y" "r \ rai_ub y" then have y [simp]: "y = (p,l,r)" by (cases y, auto) from rai have ur: "unique_root y" by auto show ?thesis proof (cases "x < l \ x > r") case True { assume xl: "x < l" hence "real_of_rat x < of_rat l" unfolding of_rat_less by auto with rai have "of_rat x < the_unique_root y" by (auto elim!: invariant_1E) with xl rai have ?thesis by (cases y, auto simp: compare_real_def comparator_of_def) } moreover { assume xr: "\ x < l" "x > r" hence "real_of_rat x > of_rat r" unfolding of_rat_less by auto with rai have "of_rat x > the_unique_root y" by (auto elim!: invariant_1E) with xr rai have ?thesis by (cases y, auto simp: compare_real_def comparator_of_def) } ultimately show ?thesis using True by auto next case False have 0: "ipoly p (real_of_rat x) \ 0" by (rule poly_cond2_no_rat_root, insert rai, auto) with rai have diff: "real_of_1 y \ of_rat x" by (auto elim!: invariant_1E) have "\ P. (1 < degree (poly_real_alg_1 y) \ \!x. root_cond y x \ poly_cond p \ P) \ P" using poly_real_alg_1.simps y rai invariant_1_2E invariant_1E by metis from this[OF gt_rat_sign_change] False have left: "compare_rat_1 x y = (if real_of_rat x \ the_unique_root y then Lt else Gt)" by (auto simp:poly_cond2_def) also have "\ = compare (real_of_rat x) (real_of_1 y)" using diff by (auto simp: compare_real_def comparator_of_def) finally show ?thesis . qed qed lemma cf_pos_0[simp]: "\ cf_pos 0" unfolding cf_pos_def by auto (* ********************* *) subsubsection\Negation\ fun uminus_1 :: "real_alg_1 \ real_alg_1" where "uminus_1 (p,l,r) = (abs_int_poly (poly_uminus p), -r, -l)" lemma uminus_1: assumes x: "invariant_1 x" defines y: "y \ uminus_1 x" shows "invariant_1 y \ (real_of_1 y = - real_of_1 x)" proof (cases x) case plr: (fields p l r) from x plr have inv: "invariant_1 (p,l,r)" by auto note * = invariant_1D[OF this] from plr have x: "x = (p,l,r)" by simp let ?p = "poly_uminus p" let ?mp = "abs_int_poly ?p" have y: "y = (?mp, -r , -l)" unfolding y plr by (simp add: Let_def) { fix y assume "root_cond (?mp, - r, - l) y" hence mpy: "ipoly ?mp y = 0" and bnd: "- of_rat r \ y" "y \ - of_rat l" unfolding root_cond_def by (auto simp: of_rat_minus) from mpy have id: "ipoly p (- y) = 0" by auto from bnd have bnd: "of_rat l \ - y" "-y \ of_rat r" by auto from id bnd have "root_cond (p, l, r) (-y)" unfolding root_cond_def by auto with inv x have "real_of_1 x = -y" by (auto intro!: the_unique_root_eqI) then have "-real_of_1 x = y" by auto } note inj = this have rc: "root_cond (?mp, - r, - l) (- real_of_1 x)" using * unfolding root_cond_def y x by (auto simp: of_rat_minus sgn_minus_rat) from inj rc have ur': "unique_root (?mp, -r, -l)" by (auto intro: unique_rootI) with rc have the: "- real_of_1 x = the_unique_root (?mp, -r, -l)" by (auto intro: the_unique_root_eqI) have xp: "p represents (real_of_1 x)" using * unfolding root_cond_def split represents_def x by auto from * have mon: "lead_coeff ?mp > 0" by (unfold pos_poly_abs_poly, auto) from poly_uminus_irreducible * have mi: "irreducible ?mp" by auto from mi mon have pc': "poly_cond ?mp" by (auto simp: cf_pos_def) from poly_condD[OF pc'] have irr: "irreducible ?mp" by auto show ?thesis unfolding y apply (intro invariant_1_realI ur' rc) using pc' inv by auto qed lemma uminus_1_2: assumes x: "invariant_1_2 x" defines y: "y \ uminus_1 x" shows "invariant_1_2 y \ (real_of_1 y = - real_of_1 x)" proof - from x have "invariant_1 x" by auto from uminus_1[OF this] have *: "real_of_1 y = - real_of_1 x" "invariant_1 y" unfolding y by auto obtain p l r where id: "x = (p,l,r)" by (cases x) from x[unfolded id] have "degree p > 1" by auto moreover have "poly_real_alg_1 y = abs_int_poly (poly_uminus p)" unfolding y id uminus_1.simps split Let_def by auto ultimately have "degree (poly_real_alg_1 y) > 1" by simp with * show ?thesis by auto qed fun uminus_2 :: "real_alg_2 \ real_alg_2" where "uminus_2 (Rational r) = Rational (-r)" | "uminus_2 (Irrational n x) = real_alg_2 (uminus_1 x)" lemma uminus_2: assumes "invariant_2 x" shows "real_of_2 (uminus_2 x) = uminus (real_of_2 x)" "invariant_2 (uminus_2 x)" using assms real_alg_2 uminus_1 by (atomize(full), cases x, auto simp: hom_distribs) declare uminus_1.simps[simp del] lift_definition uminus_3 :: "real_alg_3 \ real_alg_3" is uminus_2 by (auto simp: uminus_2) lemma uminus_3: "real_of_3 (uminus_3 x) = - real_of_3 x" by (transfer, auto simp: uminus_2) instantiation real_alg :: uminus begin lift_definition uminus_real_alg :: "real_alg \ real_alg" is uminus_3 by (simp add: uminus_3) instance .. end lemma uminus_real_alg: "- (real_of x) = real_of (- x)" by (transfer, rule uminus_3[symmetric]) (* ********************* *) subsubsection\Inverse\ fun inverse_1 :: "real_alg_1 \ real_alg_2" where "inverse_1 (p,l,r) = real_alg_2 (abs_int_poly (reflect_poly p), inverse r, inverse l)" lemma invariant_1_2_of_rat: assumes rc: "invariant_1_2 rai" shows "real_of_1 rai \ of_rat x" proof - obtain p l r where rai: "rai = (p, l, r)" by (cases rai, auto) from rc[unfolded rai] have "poly_cond2 p" "ipoly p (the_unique_root (p, l, r)) = 0" by (auto elim!: invariant_1E) from poly_cond2_no_rat_root[OF this(1), of x] this(2) show ?thesis unfolding rai by auto qed lemma inverse_1: assumes rcx: "invariant_1_2 x" defines y: "y \ inverse_1 x" shows "invariant_2 y \ (real_of_2 y = inverse (real_of_1 x))" proof (cases x) case x: (fields p l r) from x rcx have rcx: "invariant_1_2 (p,l,r)" by auto from invariant_1_2_poly_cond2[OF rcx] have pc2: "poly_cond2 p" by simp have x0: "real_of_1 (p,l,r) \ 0" using invariant_1_2_of_rat[OF rcx, of 0] x by auto let ?x = "real_of_1 (p,l,r)" let ?mp = "abs_int_poly (reflect_poly p)" from x0 rcx have lr0: "l \ 0" and "r \ 0" by auto from x0 rcx have y: "y = real_alg_2 (?mp, inverse r, inverse l)" unfolding y x Let_def inverse_1.simps by auto from rcx have mon: "lead_coeff ?mp > 0" by (unfold lead_coeff_abs_int_poly, auto) { fix y assume "root_cond (?mp, inverse r, inverse l) y" hence mpy: "ipoly ?mp y = 0" and bnd: "inverse (of_rat r) \ y" "y \ inverse (of_rat l)" unfolding root_cond_def by (auto simp: of_rat_inverse) from sgn_real_mono[OF bnd(1)] sgn_real_mono[OF bnd(2)] have "sgn (of_rat r) \ sgn y" "sgn y \ sgn (of_rat l)" by (simp_all add: algebra_simps) with rcx have sgn: "sgn (inverse (of_rat r)) = sgn y" "sgn y = sgn (inverse (of_rat l))" unfolding sgn_inverse inverse_sgn by (auto simp add: real_of_rat_sgn intro: order_antisym) from sgn[simplified, unfolded real_of_rat_sgn] lr0 have "y \ 0" by (auto simp:sgn_0_0) with mpy have id: "ipoly p (inverse y) = 0" by (auto simp: ipoly_reflect_poly) from inverse_le_sgn[OF sgn(1) bnd(1)] inverse_le_sgn[OF sgn(2) bnd(2)] have bnd: "of_rat l \ inverse y" "inverse y \ of_rat r" by auto from id bnd have "root_cond (p,l,r) (inverse y)" unfolding root_cond_def by auto from rcx this x0 have "?x = inverse y" by auto then have "inverse ?x = y" by auto } note inj = this have rc: "root_cond (?mp, inverse r, inverse l) (inverse ?x)" using rcx x0 apply (elim invariant_1_2E invariant_1E) by (simp add: root_cond_def of_rat_inverse real_of_rat_sgn inverse_le_iff_sgn ipoly_reflect_poly) from inj rc have ur: "unique_root (?mp, inverse r, inverse l)" by (auto intro: unique_rootI) with rc have the: "the_unique_root (?mp, inverse r, inverse l) = inverse ?x" by (auto intro: the_unique_root_eqI) have xp: "p represents ?x" unfolding split represents_def using rcx by (auto elim!: invariant_1E) from reflect_poly_irreducible[OF _ xp x0] poly_condD rcx have mi: "irreducible ?mp" by auto from mi mon have un: "poly_cond ?mp" by (auto simp: poly_cond_def) show ?thesis using rcx rc ur unfolding y by (intro invariant_2_realI, auto simp: x y un) qed fun inverse_2 :: "real_alg_2 \ real_alg_2" where "inverse_2 (Rational r) = Rational (inverse r)" | "inverse_2 (Irrational n x) = inverse_1 x" lemma inverse_2: assumes "invariant_2 x" shows "real_of_2 (inverse_2 x) = inverse (real_of_2 x)" "invariant_2 (inverse_2 x)" using assms by (atomize(full), cases x, auto simp: real_alg_2 inverse_1 hom_distribs) lift_definition inverse_3 :: "real_alg_3 \ real_alg_3" is inverse_2 by (auto simp: inverse_2) lemma inverse_3: "real_of_3 (inverse_3 x) = inverse (real_of_3 x)" by (transfer, auto simp: inverse_2) (* ********************* *) subsubsection\Floor\ fun floor_1 :: "real_alg_1 \ int" where "floor_1 (p,l,r) = (let (l',r',sr') = tighten_poly_bounds_epsilon p (1/2) l r (sgn (ipoly p r)); fr = floor r'; fl = floor l'; fr' = rat_of_int fr in (if fr = fl then fr else let (l'',r'',sr'') = tighten_poly_bounds_for_x p fr' l' r' sr' in if fr' < l'' then fr else fl))" lemma floor_1: assumes "invariant_1_2 x" shows "floor (real_of_1 x) = floor_1 x" proof (cases x) case (fields p l r) obtain l' r' sr' where tbe: "tighten_poly_bounds_epsilon p (1 / 2) l r (sgn (ipoly p r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?fr = "floor r'" let ?fl = "floor l'" let ?fr' = "rat_of_int ?fr" obtain l'' r'' sr'' where tbx: "tighten_poly_bounds_for_x p ?fr' l' r' sr' = (l'',r'',sr'')" by (cases rule: prod_cases3, auto) note rc = assms[unfolded fields] hence rc1: "invariant_1 (p,l,r)" by auto have id: "floor_1 x = ((if ?fr = ?fl then ?fr else if ?fr' < l'' then ?fr else ?fl))" unfolding fields floor_1.simps tbe Let_def split tbx by simp let ?x = "real_of_1 x" have x: "?x = the_unique_root (p,l,r)" unfolding fields by simp have bnd: "l \ l'" "r' \ r" "r' - l' \ 1 / 2" and rc': "root_cond (p, l', r') (the_unique_root (p, l, r))" and sr': "sr' = sgn (ipoly p r')" by (atomize(full), intro conjI tighten_poly_bounds_epsilon[OF _ _ tbe refl],insert rc,auto elim!: invariant_1E) let ?r = real_of_rat from rc'[folded x, unfolded split] have ineq: "?r l' \ ?x" "?x \ ?r r'" "?r l' \ ?r r'" by auto hence lr': "l' \ r'" unfolding of_rat_less_eq by simp have flr: "?fl \ ?fr" by (rule floor_mono[OF lr']) from invariant_1_sub_interval[OF rc1 rc' bnd(1,2)] have rc': "invariant_1 (p, l', r')" and id': "the_unique_root (p, l', r') = the_unique_root (p, l, r)" by auto with rc have rc2': "invariant_1_2 (p, l', r')" by auto have x: "?x = the_unique_root (p,l',r')" unfolding fields using id' by simp { assume "?fr \ ?fl" with flr have flr: "?fl \ ?fr - 1" by simp have "?fr' \ r'" "l' \ ?fr'" using flr bnd by linarith+ } note fl_diff = this show ?thesis proof (cases "?fr = ?fl") case True hence id1: "floor_1 x = ?fr" unfolding id by auto from True have id: "floor (?r l') = floor (?r r')" by simp have "floor ?x \ floor (?r r')" by (rule floor_mono[OF ineq(2)]) moreover have "floor (?r l') \ floor ?x" by (rule floor_mono[OF ineq(1)]) ultimately have "floor ?x = floor (?r r')" unfolding id by (simp add: id) then show ?thesis by (simp add: id1) next case False with id have id: "floor_1 x = (if ?fr' < l'' then ?fr else ?fl)" by simp from rc2' have "unique_root (p,l',r')" "poly_cond2 p" by auto from tighten_poly_bounds_for_x[OF this tbx sr'] have ineq': "l' \ l''" "r'' \ r'" and lr'': "l'' \ r''" and rc'': "root_cond (p,l'',r'') ?x" and fr': "\ (l'' \ ?fr' \ ?fr' \ r'')" unfolding x by auto from rc''[unfolded split] have ineq'': "?r l'' \ ?x" "?x \ ?r r''" by auto from False have "?fr \ ?fl" by auto note fr = fl_diff[OF this] show ?thesis proof (cases "?fr' < l''") case True with id have id: "floor_1 x = ?fr" by simp have "floor ?x \ ?fr" using floor_mono[OF ineq(2)] by simp moreover from True have "?r ?fr' < ?r l''" unfolding of_rat_less . with ineq''(1) have "?r ?fr' \ ?x" by simp from floor_mono[OF this] have "?fr \ floor ?x" by simp ultimately show ?thesis unfolding id by auto next case False with id have id: "floor_1 x = ?fl" by simp from False have "l'' \ ?fr'" by auto from floor_mono[OF ineq(1)] have "?fl \ floor ?x" by simp moreover have "floor ?x \ ?fl" proof - from False fr' have fr': "r'' < ?fr'" by auto hence "floor r'' < ?fr" by linarith with floor_mono[OF ineq''(2)] have "floor ?x \ ?fr - 1" by auto also have "?fr - 1 = floor (r' - 1)" by simp also have "\ \ ?fl" by (rule floor_mono, insert bnd, auto) finally show ?thesis . qed ultimately show ?thesis unfolding id by auto qed qed qed (* ********************* *) subsubsection\Generic Factorization and Bisection Framework\ lemma card_1_Collect_ex1: assumes "card (Collect P) = 1" shows "\! x. P x" proof - from assms[unfolded card_eq_1_iff] obtain x where "Collect P = {x}" by auto thus ?thesis by (intro ex1I[of _ x], auto) qed fun sub_interval :: "rat \ rat \ rat \ rat \ bool" where "sub_interval (l,r) (l',r') = (l' \ l \ r \ r')" fun in_interval :: "rat \ rat \ real \ bool" where "in_interval (l,r) x = (of_rat l \ x \ x \ of_rat r)" definition converges_to :: "(nat \ rat \ rat) \ real \ bool" where "converges_to f x \ (\ n. in_interval (f n) x \ sub_interval (f (Suc n)) (f n)) \ (\ (eps :: real) > 0. \ n l r. f n = (l,r) \ of_rat r - of_rat l \ eps)" context fixes bnd_update :: "'a \ 'a" and bnd_get :: "'a \ rat \ rat" begin definition at_step :: "(nat \ rat \ rat) \ nat \ 'a \ bool" where "at_step f n a \ \ i. bnd_get ((bnd_update ^^ i) a) = f (n + i)" partial_function (tailrec) select_correct_factor_main :: "'a \ (int poly \ root_info)list \ (int poly \ root_info)list \ rat \ rat \ nat \ (int poly \ root_info) \ rat \ rat" where [code]: "select_correct_factor_main bnd todo old l r n = (case todo of Nil \ if n = 1 then (hd old, l, r) else let bnd' = bnd_update bnd in (case bnd_get bnd' of (l,r) \ select_correct_factor_main bnd' old [] l r 0) | Cons (p,ri) todo \ let m = root_info.l_r ri l r in if m = 0 then select_correct_factor_main bnd todo old l r n else select_correct_factor_main bnd todo ((p,ri) # old) l r (n + m))" definition select_correct_factor :: "'a \ (int poly \ root_info)list \ (int poly \ root_info) \ rat \ rat" where "select_correct_factor init polys = (case bnd_get init of (l,r) \ select_correct_factor_main init polys [] l r 0)" lemma select_correct_factor_main: assumes conv: "converges_to f x" and at: "at_step f i a" and res: "select_correct_factor_main a todo old l r n = ((q,ri_fin),(l_fin,r_fin))" and bnd: "bnd_get a = (l,r)" and ri: "\ q ri. (q,ri) \ set todo \ set old \ root_info_cond ri q" and q0: "\ q ri. (q,ri) \ set todo \ set old \ q \ 0" and ex: "\q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0" and dist: "distinct (map fst (todo @ old))" and old: "\ q ri. (q,ri) \ set old \ root_info.l_r ri l r \ 0" and un: "\ x :: real. (\q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0) \ \!q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0" and n: "n = sum_list (map (\ (q,ri). root_info.l_r ri l r) old)" shows "unique_root (q,l_fin,r_fin) \ (q,ri_fin) \ set todo \ set old \ x = the_unique_root (q,l_fin,r_fin)" proof - define orig where "orig = set todo \ set old" have orig: "set todo \ set old \ orig" unfolding orig_def by auto let ?rts = "{x :: real. \ q ri. (q,ri) \ orig \ ipoly q x = 0}" define rts where "rts = ?rts" let ?h = "\ (x,y). abs (x - y)" let ?r = real_of_rat have rts: "?rts = (\ ((\ (q,ri). {x. ipoly q x = 0}) ` set (todo @ old)))" unfolding orig_def by auto have "finite rts" unfolding rts rts_def using finite_ipoly_roots[OF q0] finite_set[of "todo @ old"] by auto hence fin: "finite (rts \ rts - Id)" by auto define diffs where "diffs = insert 1 {abs (x - y) | x y. x \ rts \ y \ rts \ x \ y}" have "finite {abs (x - y) | x y. x \ rts \ y \ rts \ x \ y}" by (rule subst[of _ _ finite, OF _ finite_imageI[OF fin, of ?h]], auto) hence diffs: "finite diffs" "diffs \ {}" unfolding diffs_def by auto define eps where "eps = Min diffs / 2" have "\ x. x \ diffs \ x > 0" unfolding diffs_def by auto with Min_gr_iff[OF diffs] have eps: "eps > 0" unfolding eps_def by auto note conv = conv[unfolded converges_to_def] from conv eps obtain N L R where N: "f N = (L,R)" "?r R - ?r L \ eps" by auto obtain pair where pair: "pair = (todo,i)" by auto define rel where "rel = measures [ \ (t,i). N - i, \ (t :: (int poly \ root_info) list,i). length t]" have wf: "wf rel" unfolding rel_def by simp show ?thesis using at res bnd ri q0 ex dist old un n pair orig proof (induct pair arbitrary: todo i old a l r n rule: wf_induct[OF wf]) case (1 pair todo i old a l r n) note IH = 1(1)[rule_format] note at = 1(2) note res = 1(3)[unfolded select_correct_factor_main.simps[of _ todo]] note bnd = 1(4) note ri = 1(5) note q0 = 1(6) note ex = 1(7) note dist = 1(8) note old = 1(9) note un = 1(10) note n = 1(11) note pair = 1(12) note orig = 1(13) from at[unfolded at_step_def, rule_format, of 0] bnd have fi: "f i = (l,r)" by auto with conv have inx: "in_interval (f i) x" by blast hence lxr: "?r l \ x" "x \ ?r r" unfolding fi by auto from order.trans[OF this] have lr: "l \ r" unfolding of_rat_less_eq . show ?case proof (cases todo) case (Cons rri tod) obtain s ri where rri: "rri = (s,ri)" by force with Cons have todo: "todo = (s,ri) # tod" by simp note res = res[unfolded todo list.simps split Let_def] from root_info_condD(1)[OF ri[of s ri, unfolded todo] lr] have ri': "root_info.l_r ri l r = card {x. root_cond (s, l, r) x}" by auto from q0 have s0: "s \ 0" unfolding todo by auto from finite_ipoly_roots[OF s0] have fins: "finite {x. root_cond (s, l, r) x}" unfolding root_cond_def by auto have rel: "((tod,i), pair) \ rel" unfolding rel_def pair todo by simp show ?thesis proof (cases "root_info.l_r ri l r = 0") case True with res have res: "select_correct_factor_main a tod old l r n = ((q, ri_fin), l_fin, r_fin)" by auto from ri'[symmetric, unfolded True] fins have empty: "{x. root_cond (s, l, r) x} = {}" by simp from ex lxr empty have ex': "(\q. q \ fst ` set tod \ fst ` set old \ ipoly q x = 0)" unfolding todo root_cond_def split by auto have "unique_root (q, l_fin, r_fin) \ (q, ri_fin) \ set tod \ set old \ x = the_unique_root (q, l_fin, r_fin)" proof (rule IH[OF rel at res bnd ri _ ex' _ _ _ n refl], goal_cases) case (5 y) thus ?case using un[of y] unfolding todo by auto next case 2 thus ?case using q0 unfolding todo by auto qed (insert dist old orig, auto simp: todo) thus ?thesis unfolding todo by auto next case False with res have res: "select_correct_factor_main a tod ((s, ri) # old) l r (n + root_info.l_r ri l r) = ((q, ri_fin), l_fin, r_fin)" by auto from ex have ex': "\q. q \ fst ` set tod \ fst ` set ((s, ri) # old) \ ipoly q x = 0" unfolding todo by auto from dist have dist: "distinct (map fst (tod @ (s, ri) # old))" unfolding todo by auto have id: "set todo \ set old = set tod \ set ((s, ri) # old)" unfolding todo by simp show ?thesis unfolding id proof (rule IH[OF rel at res bnd ri _ ex' dist], goal_cases) case 4 thus ?case using un unfolding todo by auto qed (insert old False orig, auto simp: q0 todo n) qed next case Nil note res = res[unfolded Nil list.simps Let_def] from ex[unfolded Nil] lxr obtain s where "s \ fst ` set old \ root_cond (s,l,r) x" unfolding root_cond_def by auto then obtain q1 ri1 old' where old': "old = (q1,ri1) # old'" using id by (cases old, auto) let ?ri = "root_info.l_r ri1 l r" from old[unfolded old'] have 0: "?ri \ 0" by auto from n[unfolded old'] 0 have n0: "n \ 0" by auto from ri[unfolded old'] have ri': "root_info_cond ri1 q1" by auto show ?thesis proof (cases "n = 1") case False with n0 have n1: "n > 1" by auto obtain l' r' where bnd': "bnd_get (bnd_update a) = (l',r')" by force with res False have res: "select_correct_factor_main (bnd_update a) old [] l' r' 0 = ((q, ri_fin), l_fin, r_fin)" by auto have at': "at_step f (Suc i) (bnd_update a)" unfolding at_step_def proof (intro allI, goal_cases) case (1 n) have id: "(bnd_update ^^ Suc n) a = (bnd_update ^^ n) (bnd_update a)" by (induct n, auto) from at[unfolded at_step_def, rule_format, of "Suc n"] show ?case unfolding id by simp qed from 0[unfolded root_info_condD(1)[OF ri' lr]] obtain y1 where y1: "root_cond (q1,l,r) y1" by (cases "Collect (root_cond (q1, l, r)) = {}", auto) from n1[unfolded n old'] have "?ri > 1 \ sum_list (map (\ (q,ri). root_info.l_r ri l r) old') \ 0" by (cases "sum_list (map (\ (q,ri). root_info.l_r ri l r) old')", auto) hence "\ q2 ri2 y2. (q2,ri2) \ set old \ root_cond (q2,l,r) y2 \ y1 \ y2" proof assume "?ri > 1" with root_info_condD(1)[OF ri' lr] have "card {x. root_cond (q1, l, r) x} > 1" by simp from card_gt_1D[OF this] y1 obtain y2 where "root_cond (q1,l,r) y2" and "y1 \ y2" by auto thus ?thesis unfolding old' by auto next assume "sum_list (map (\ (q,ri). root_info.l_r ri l r) old') \ 0" then obtain q2 ri2 where mem: "(q2,ri2) \ set old'" and ri2: "root_info.l_r ri2 l r \ 0" by auto with q0 ri have "root_info_cond ri2 q2" unfolding old' by auto from ri2[unfolded root_info_condD(1)[OF this lr]] obtain y2 where y2: "root_cond (q2,l,r) y2" by (cases "Collect (root_cond (q2, l, r)) = {}", auto) from dist[unfolded old'] split_list[OF mem] have diff: "q1 \ q2" by auto from y1 have q1: "q1 \ fst ` set todo \ fst ` set old \ ipoly q1 y1 = 0" unfolding old' root_cond_def by auto from y2 have q2: "q2 \ fst ` set todo \ fst ` set old \ ipoly q2 y2 = 0" unfolding old' root_cond_def using mem by force have "y1 \ y2" proof assume id: "y1 = y2" from q1 have "\ q1. q1 \ fst ` set todo \ fst ` set old \ ipoly q1 y1 = 0" by blast from un[OF this] q1 q2[folded id] have "q1 = q2" by auto with diff show False by simp qed with mem y2 show ?thesis unfolding old' by auto qed then obtain q2 ri2 y2 where mem2: "(q2,ri2) \ set old" and y2: "root_cond (q2,l,r) y2" and diff: "y1 \ y2" by auto from mem2 orig have "(q1,ri1) \ orig" "(q2,ri2) \ orig" unfolding old' by auto with y1 y2 diff have "abs (y1 - y2) \ diffs" unfolding diffs_def rts_def root_cond_def by auto from Min_le[OF diffs(1) this] have "abs (y1 - y2) \ 2 * eps" unfolding eps_def by auto with eps have eps: "abs (y1 - y2) > eps" by auto from y1 y2 have l: "of_rat l \ min y1 y2" unfolding root_cond_def by auto from y1 y2 have r: "of_rat r \ max y1 y2" unfolding root_cond_def by auto from l r eps have eps: "of_rat r - of_rat l > eps" by auto have "i < N" proof (rule ccontr) assume "\ i < N" hence "\ k. i = N + k" by presburger then obtain k where i: "i = N + k" by auto { fix k l r assume "f (N + k) = (l,r)" hence "of_rat r - of_rat l \ eps" proof (induct k arbitrary: l r) case 0 with N show ?case by auto next case (Suc k l r) obtain l' r' where f: "f (N + k) = (l',r')" by force from Suc(1)[OF this] have IH: "?r r' - ?r l' \ eps" by auto from f Suc(2) conv[THEN conjunct1, rule_format, of "N + k"] have "?r l \ ?r l'" "?r r \ ?r r'" by (auto simp: of_rat_less_eq) thus ?case using IH by auto qed } note * = this from at[unfolded at_step_def i, rule_format, of 0] bnd have "f (N + k) = (l,r)" by auto from *[OF this] eps show False by auto qed hence rel: "((old, Suc i), pair) \ rel" unfolding pair rel_def by auto from dist have dist: "distinct (map fst (old @ []))" unfolding Nil by auto have id: "set todo \ set old = set old \ set []" unfolding Nil by auto show ?thesis unfolding id proof (rule IH[OF rel at' res bnd' ri _ _ dist _ _ _ refl], goal_cases) case 2 thus ?case using q0 by auto qed (insert ex un orig Nil, auto) next case True with res old' have id: "q = q1" "ri_fin = ri1" "l_fin = l" "r_fin = r" by auto from n[unfolded True old'] 0 have 1: "?ri = 1" by (cases ?ri; cases "?ri - 1", auto) from root_info_condD(1)[OF ri' lr] 1 have "card {x. root_cond (q1,l,r) x} = 1" by auto from card_1_Collect_ex1[OF this] have unique: "unique_root (q1,l,r)" . from ex[unfolded Nil old'] consider (A) "ipoly q1 x = 0" | (B) q where "q \ fst ` set old'" "ipoly q x = 0" by auto hence "x = the_unique_root (q1,l,r)" proof (cases) case A with lxr have "root_cond (q1,l,r) x" unfolding root_cond_def by auto from the_unique_root_eqI[OF unique this] show ?thesis by simp next case (B q) with lxr have "root_cond (q,l,r) x" unfolding root_cond_def by auto hence empty: "{x. root_cond (q,l,r) x} \ {}" by auto from B(1) obtain ri' where mem: "(q,ri') \ set old'" by force from q0[unfolded old'] mem have q0: "q \ 0" by auto from finite_ipoly_roots[OF this] have "finite {x. root_cond (q,l,r) x}" unfolding root_cond_def by auto with empty have card: "card {x. root_cond (q,l,r) x} \ 0" by simp from ri[unfolded old'] mem have "root_info_cond ri' q" by auto from root_info_condD(1)[OF this lr] card have "root_info.l_r ri' l r \ 0" by auto with n[unfolded True old'] 1 split_list[OF mem] have False by auto thus ?thesis by simp qed thus ?thesis unfolding id using unique ri' unfolding old' by auto qed qed qed qed lemma select_correct_factor: assumes conv: "converges_to (\ i. bnd_get ((bnd_update ^^ i) init)) x" and res: "select_correct_factor init polys = ((q,ri),(l,r))" and ri: "\ q ri. (q,ri) \ set polys \ root_info_cond ri q" and q0: "\ q ri. (q,ri) \ set polys \ q \ 0" and ex: "\q. q \ fst ` set polys \ ipoly q x = 0" and dist: "distinct (map fst polys)" and un: "\ x :: real. (\q. q \ fst ` set polys \ ipoly q x = 0) \ \!q. q \ fst ` set polys \ ipoly q x = 0" shows "unique_root (q,l,r) \ (q,ri) \ set polys \ x = the_unique_root (q,l,r)" proof - obtain l' r' where init: "bnd_get init = (l',r')" by force from res[unfolded select_correct_factor_def init split] have res: "select_correct_factor_main init polys [] l' r' 0 = ((q, ri), l, r)" by auto have at: "at_step (\ i. bnd_get ((bnd_update ^^ i) init)) 0 init" unfolding at_step_def by auto have "unique_root (q,l,r) \ (q,ri) \ set polys \ set [] \ x = the_unique_root (q,l,r)" by (rule select_correct_factor_main[OF conv at res init ri], insert dist un ex q0, auto) thus ?thesis by auto qed definition real_alg_2' :: "root_info \ int poly \ rat \ rat \ real_alg_2" where [code del]: "real_alg_2' ri p l r = ( if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else real_alg_2_main ri (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l',r',sr') \ (p, l', r')))" lemma real_alg_2'_code[code]: "real_alg_2' ri p l r = (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else case normalize_bounds_1 (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l', r', sr') \ (p, l', r')) of (p', l, r) \ Irrational (root_info.number_root ri r) (p', l, r))" unfolding real_alg_2'_def real_alg_2_main_def by (cases "tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r))", simp add: Let_def) definition real_alg_2'' :: "root_info \ int poly \ rat \ rat \ real_alg_2" where "real_alg_2'' ri p l r = (case normalize_bounds_1 (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l', r', sr') \ (p, l', r')) of (p', l, r) \ Irrational (root_info.number_root ri r) (p', l, r))" lemma real_alg_2'': "degree p \ 1 \ real_alg_2'' ri p l r = real_alg_2' ri p l r" unfolding real_alg_2'_code real_alg_2''_def by auto lemma poly_cond_degree_0_imp_no_root: fixes x :: "'b :: {comm_ring_1,ring_char_0}" assumes pc: "poly_cond p" and deg: "degree p = 0" shows "ipoly p x \ 0" proof from pc have "p \ 0" by auto moreover assume "ipoly p x = 0" note poly_zero[OF this] ultimately show False using deg by auto qed lemma real_alg_2': assumes ur: "unique_root (q,l,r)" and pc: "poly_cond q" and ri: "root_info_cond ri q" shows "invariant_2 (real_alg_2' ri q l r) \ real_of_2 (real_alg_2' ri q l r) = the_unique_root (q,l,r)" (is "_ \ _ = ?x") proof (cases "degree q" "Suc 0" rule: linorder_cases) case deg: less then have "degree q = 0" by auto from poly_cond_degree_0_imp_no_root[OF pc this] ur have False by force then show ?thesis by auto next case deg: equal hence id: "real_alg_2' ri q l r = Rational (Rat.Fract (- coeff q 0) (coeff q 1))" unfolding real_alg_2'_def by auto show ?thesis unfolding id using degree_1_ipoly[OF deg] using unique_rootD(4)[OF ur] by auto next case deg: greater with pc have pc2: "poly_cond2 q" by auto let ?rai = "real_alg_2' ri q l r" let ?r = real_of_rat obtain l' r' sr' where tight: "tighten_poly_bounds_for_x q 0 l r (sgn (ipoly q r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?rai' = "(q, l', r')" have rai': "?rai = real_alg_2_main ri ?rai'" unfolding real_alg_2'_def using deg tight by auto hence rai: "real_of_1 ?rai' = the_unique_root (q,l',r')" by auto note tight = tighten_poly_bounds_for_x[OF ur pc2 tight refl] let ?x = "the_unique_root (q, l, r)" from tight have tight: "root_cond (q,l',r') ?x" "l \ l'" "l' \ r'" "r' \ r" "l' > 0 \ r' < 0" by auto from unique_root_sub_interval[OF ur tight(1) tight(2,4)] poly_condD[OF pc] have ur': "unique_root (q, l', r')" and x: "?x = the_unique_root (q,l',r')" by auto from tight(2-) have sgn: "sgn l' = sgn r'" by auto show ?thesis unfolding rai' using real_alg_2_main[of ?rai' ri] invariant_1_realI[of ?rai' ?x] by (auto simp: tight(1) sgn pc ri ur') qed definition select_correct_factor_int_poly :: "'a \ int poly \ real_alg_2" where "select_correct_factor_int_poly init p \ let qs = factors_of_int_poly p; polys = map (\ q. (q, root_info q)) qs; ((q,ri),(l,r)) = select_correct_factor init polys in real_alg_2' ri q l r" lemma select_correct_factor_int_poly: assumes conv: "converges_to (\ i. bnd_get ((bnd_update ^^ i) init)) x" and rai: "select_correct_factor_int_poly init p = rai" and x: "ipoly p x = 0" and p: "p \ 0" shows "invariant_2 rai \ real_of_2 rai = x" proof - obtain qs where fact: "factors_of_int_poly p = qs" by auto define polys where "polys = map (\ q. (q, root_info q)) qs" obtain q ri l r where res: "select_correct_factor init polys = ((q,ri),(l,r))" by (cases "select_correct_factor init polys", auto) have fst: "map fst polys = qs" "fst ` set polys = set qs" unfolding polys_def map_map o_def by force+ note fact' = factors_of_int_poly[OF fact] note rai = rai[unfolded select_correct_factor_int_poly_def Let_def fact, folded polys_def, unfolded res split] from fact' fst have dist: "distinct (map fst polys)" by auto from fact'(2)[OF p, of x] x fst have ex: "\q. q \ fst ` set polys \ ipoly q x = 0" by auto { fix q ri assume "(q,ri) \ set polys" hence ri: "ri = root_info q" and q: "q \ set qs" unfolding polys_def by auto from fact'(1)[OF q] have *: "lead_coeff q > 0" "irreducible q" "degree q > 0" by auto from * have q0: "q \ 0" by auto from root_info[OF *(2-3)] ri have ri: "root_info_cond ri q" by auto note ri q0 * } note polys = this have "unique_root (q, l, r) \ (q, ri) \ set polys \ x = the_unique_root (q, l, r)" by (rule select_correct_factor[OF conv res polys(1) _ ex dist, unfolded fst, OF _ _ fact'(3)[OF p]], insert fact'(2)[OF p] polys(2), auto) hence ur: "unique_root (q,l,r)" and mem: "(q,ri) \ set polys" and x: "x = the_unique_root (q,l,r)" by auto note polys = polys[OF mem] from polys(3-4) have ty: "poly_cond q" by (simp add: poly_cond_def) show ?thesis unfolding x rai[symmetric] by (intro real_alg_2' ur ty polys(1)) qed end (* ********************* *) subsubsection\Addition\ lemma ipoly_0_0[simp]: "ipoly f (0::'a::{comm_ring_1,ring_char_0}) = 0 \ poly f 0 = 0" unfolding poly_0_coeff_0 by simp lemma add_rat_roots_below[simp]: "roots_below (poly_add_rat r p) x = (\y. y + of_rat r) ` roots_below p (x - of_rat r)" proof (unfold add_rat_roots image_def, intro Collect_eqI, goal_cases) case (1 y) then show ?case by (auto intro: exI[of _ "y - real_of_rat r"]) qed lemma add_rat_root_cond: shows "root_cond (cf_pos_poly (poly_add_rat m p),l,r) x = root_cond (p, l - m, r - m) (x - of_rat m)" by (unfold root_cond_def, auto simp add: add_rat_roots hom_distribs) lemma add_rat_unique_root: "unique_root (cf_pos_poly (poly_add_rat m p), l, r) = unique_root (p, l-m, r-m)" by (auto simp: add_rat_root_cond) fun add_rat_1 :: "rat \ real_alg_1 \ real_alg_1" where "add_rat_1 r1 (p2,l2,r2) = ( let p = cf_pos_poly (poly_add_rat r1 p2); (l,r,sr) = tighten_poly_bounds_for_x p 0 (l2+r1) (r2+r1) (sgn (ipoly p (r2+r1))) in (p,l,r))" lemma poly_real_alg_1_add_rat[simp]: "poly_real_alg_1 (add_rat_1 r y) = cf_pos_poly (poly_add_rat r (poly_real_alg_1 y))" by (cases y, auto simp: Let_def split: prod.split) lemma sgn_cf_pos: assumes "lead_coeff p > 0" shows "sgn (ipoly (cf_pos_poly p) (x::'a::linordered_field)) = sgn (ipoly p x)" proof (cases "p = 0") case True with assms show ?thesis by auto next case False from cf_pos_poly_main False obtain d where p': "Polynomial.smult d (cf_pos_poly p) = p" by auto have "d > 0" proof (rule zero_less_mult_pos2) from False assms have "0 < lead_coeff p" by (auto simp: cf_pos_def) also from p' have "\ = d * lead_coeff (cf_pos_poly p)" by (metis lead_coeff_smult) finally show "0 < \". show "lead_coeff (cf_pos_poly p) > 0" using False by (unfold lead_coeff_cf_pos_poly) qed moreover from p' have "ipoly p x = of_int d * ipoly (cf_pos_poly p) x" by (fold poly_smult of_int_hom.map_poly_hom_smult, auto) ultimately show ?thesis by (auto simp: sgn_mult[where 'a='a]) qed lemma add_rat_1: fixes r1 :: rat assumes inv_y: "invariant_1_2 y" defines "z \ add_rat_1 r1 y" shows "invariant_1_2 z \ (real_of_1 z = of_rat r1 + real_of_1 y)" proof (cases y) case y_def: (fields p2 l2 r2) define p where "p \ cf_pos_poly (poly_add_rat r1 p2)" obtain l r sr where lr: "tighten_poly_bounds_for_x p 0 (l2+r1) (r2+r1) (sgn (ipoly p (r2+r1))) = (l,r,sr)" by (metis surj_pair) from lr have z: "z = (p,l,r)" by (auto simp: y_def z_def p_def Let_def) from inv_y have ur: "unique_root (p, l2 + r1, r2 + r1)" by (auto simp: p_def add_rat_root_cond y_def add_rat_unique_root) from inv_y[unfolded y_def invariant_1_2_def,simplified] have pc2: "poly_cond2 p" unfolding p_def apply (intro poly_cond2I poly_add_rat_irreducible poly_condI, unfold lead_coeff_cf_pos_poly) apply (auto elim!: invariant_1E) done note main = tighten_poly_bounds_for_x[OF ur pc2 lr refl, simplified] then have "sgn l = sgn r" unfolding sgn_if apply simp apply linarith done from invariant_1_2_realI[OF main(4) _ main(7), simplified, OF this pc2] main(1-3) ur show ?thesis by (auto simp: z p_def y_def add_rat_root_cond ex1_the_shift) qed fun tighten_poly_bounds_binary :: "int poly \ int poly \ (rat \ rat \ rat) \ rat \ rat \ rat \ (rat \ rat \ rat) \ rat \ rat \ rat" where "tighten_poly_bounds_binary cr1 cr2 ((l1,r1,sr1),(l2,r2,sr2)) = (tighten_poly_bounds cr1 l1 r1 sr1, tighten_poly_bounds cr2 l2 r2 sr2)" lemma tighten_poly_bounds_binary: assumes ur: "unique_root (p1,l1,r1)" "unique_root (p2,l2,r2)" and pt: "poly_cond2 p1" "poly_cond2 p2" defines "x \ the_unique_root (p1,l1,r1)" and "y \ the_unique_root (p2,l2,r2)" assumes bnd: "\ l1 r1 l2 r2 l r sr1 sr2. I l1 \ I l2 \ root_cond (p1,l1,r1) x \ root_cond (p2,l2,r2) y \ bnd ((l1,r1,sr1),(l2,r2,sr2)) = (l,r) \ of_rat l \ f x y \ f x y \ of_rat r" and approx: "\ l1 r1 l2 r2 l1' r1' l2' r2' l l' r r' sr1 sr2 sr1' sr2'. I l1 \ I l2 \ l1 \ r1 \ l2 \ r2 \ (l,r) = bnd ((l1,r1,sr1), (l2,r2,sr2)) \ (l',r') = bnd ((l1',r1',sr1'), (l2',r2',sr2')) \ (l1',r1') \ {(l1,(l1+r1)/2),((l1+r1)/2,r1)} \ (l2',r2') \ {(l2,(l2+r2)/2),((l2+r2)/2,r2)} \ (r' - l') \ 3/4 * (r - l) \ l \ l' \ r' \ r" and I_mono: "\ l l'. I l \ l \ l' \ I l'" and I: "I l1" "I l2" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" shows "converges_to (\ i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sr1),(l2,r2,sr2)))) (f x y)" proof - let ?upd = "tighten_poly_bounds_binary p1 p2" define upd where "upd = ?upd" define init where "init = ((l1, r1, sr1), l2, r2, sr2)" let ?g = "(\i. bnd ((upd ^^ i) init))" obtain l r where bnd_init: "bnd init = (l,r)" by force note ur1 = unique_rootD[OF ur(1)] note ur2 = unique_rootD[OF ur(2)] from ur1(4) ur2(4) x_def y_def have rc1: "root_cond (p1,l1,r1) x" and rc2: "root_cond (p2,l2,r2) y" by auto define g where "g = ?g" { fix i L1 R1 L2 R2 L R j SR1 SR2 assume "((upd ^^ i)) init = ((L1,R1,SR1),(L2,R2,SR2))" "g i = (L,R)" hence "I L1 \ I L2 \ root_cond (p1,L1,R1) x \ root_cond (p2,L2,R2) y \ unique_root (p1, L1, R1) \ unique_root (p2, L2, R2) \ in_interval (L,R) (f x y) \ (i = Suc j \ sub_interval (g i) (g j) \ (R - L \ 3/4 * (snd (g j) - fst (g j)))) \ SR1 = sgn (ipoly p1 R1) \ SR2 = sgn (ipoly p2 R2)" proof (induct i arbitrary: L1 R1 L2 R2 L R j SR1 SR2) case 0 thus ?case using I rc1 rc2 ur bnd[of l1 l2 r1 r2 sr1 sr2 L R] g_def sr unfolding init_def by auto next case (Suc i) obtain l1 r1 l2 r2 sr1 sr2 where updi: "(upd ^^ i) init = ((l1, r1, sr1), l2, r2, sr2)" by (cases "(upd ^^ i) init", auto) obtain l r where bndi: "bnd ((l1, r1, sr1), l2, r2, sr2) = (l,r)" by force hence gi: "g i = (l,r)" using updi unfolding g_def by auto have "(upd ^^ Suc i) init = upd ((l1, r1, sr1), l2, r2, sr2)" using updi by simp from Suc(2)[unfolded this] have upd: "upd ((l1, r1, sr1), l2, r2, sr2) = ((L1, R1, SR1), L2, R2, SR2)" . from upd updi Suc(3) have bndsi: "bnd ((L1, R1, SR1), L2, R2, SR2) = (L,R)" by (auto simp: g_def) from Suc(1)[OF updi gi] have I: "I l1" "I l2" and rc: "root_cond (p1,l1,r1) x" "root_cond (p2,l2,r2) y" and ur: "unique_root (p1, l1, r1)" "unique_root (p2, l2, r2)" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" by auto from upd[unfolded upd_def] have tight: "tighten_poly_bounds p1 l1 r1 sr1 = (L1, R1, SR1)" "tighten_poly_bounds p2 l2 r2 sr2 = (L2, R2, SR2)" by auto note tight1 = tighten_poly_bounds[OF tight(1) ur(1) pt(1) sr(1)] note tight2 = tighten_poly_bounds[OF tight(2) ur(2) pt(2) sr(2)] from tight1 have lr1: "l1 \ r1" by auto from tight2 have lr2: "l2 \ r2" by auto note ur1 = unique_rootD[OF ur(1)] note ur2 = unique_rootD[OF ur(2)] from tight1 I_mono[OF I(1)] have I1: "I L1" by auto from tight2 I_mono[OF I(2)] have I2: "I L2" by auto note ur1 = unique_root_sub_interval[OF ur(1) tight1(1,2,4)] note ur2 = unique_root_sub_interval[OF ur(2) tight2(1,2,4)] from rc(1) ur ur1 have x: "x = the_unique_root (p1,L1,R1)" by (auto intro!:the_unique_root_eqI) from rc(2) ur ur2 have y: "y = the_unique_root (p2,L2,R2)" by (auto intro!:the_unique_root_eqI) from unique_rootD[OF ur1(1)] x have x: "root_cond (p1,L1,R1) x" by auto from unique_rootD[OF ur2(1)] y have y: "root_cond (p2,L2,R2) y" by auto from tight(1) have half1: "(L1, R1) \ {(l1, (l1 + r1) / 2), ((l1 + r1) / 2, r1)}" unfolding tighten_poly_bounds_def Let_def by (auto split: if_splits) from tight(2) have half2: "(L2, R2) \ {(l2, (l2 + r2) / 2), ((l2 + r2) / 2, r2)}" unfolding tighten_poly_bounds_def Let_def by (auto split: if_splits) from approx[OF I lr1 lr2 bndi[symmetric] bndsi[symmetric] half1 half2] have "R - L \ 3 / 4 * (r - l) \ l \ L \ R \ r" . hence "sub_interval (g (Suc i)) (g i)" "R - L \ 3/4 * (snd (g i) - fst (g i))" unfolding gi Suc(3) by auto with bnd[OF I1 I2 x y bndsi] show ?case using I1 I2 x y ur1 ur2 tight1(6) tight2(6) by auto qed } note invariants = this define L where "L = (\ i. fst (g i))" define R where "R = (\ i. snd (g i))" { fix i obtain l1 r1 l2 r2 sr1 sr2 where updi: "(upd ^^ i) init = ((l1, r1, sr1), l2, r2, sr2)" by (cases "(upd ^^ i) init", auto) obtain l r where bnd': "bnd ((l1, r1, sr1), l2, r2, sr2) = (l,r)" by force have gi: "g i = (l,r)" unfolding g_def updi bnd' by auto hence id: "l = L i" "r = R i" unfolding L_def R_def by auto from invariants[OF updi gi[unfolded id]] have "in_interval (L i, R i) (f x y)" "\ j. i = Suc j \ sub_interval (g i) (g j) \ R i - L i \ 3 / 4 * (R j - L j)" unfolding L_def R_def by auto } note * = this { fix i from *(1)[of i] *(2)[of "Suc i", OF refl] have "in_interval (g i) (f x y)" "sub_interval (g (Suc i)) (g i)" "R (Suc i) - L (Suc i) \ 3 / 4 * (R i - L i)" unfolding L_def R_def by auto } note * = this show ?thesis unfolding upd_def[symmetric] init_def[symmetric] g_def[symmetric] unfolding converges_to_def proof (intro conjI allI impI, rule *(1), rule *(2)) fix eps :: real assume eps: "0 < eps" let ?r = real_of_rat define r where "r = (\ n. ?r (R n))" define l where "l = (\ n. ?r (L n))" define diff where "diff = (\ n. r n - l n)" { fix n from *(3)[of n] have "?r (R (Suc n) - L (Suc n)) \ ?r (3 / 4 * (R n - L n))" unfolding of_rat_less_eq by simp also have "?r (R (Suc n) - L (Suc n)) = (r (Suc n) - l (Suc n))" unfolding of_rat_diff r_def l_def by simp also have "?r (3 / 4 * (R n - L n)) = 3 / 4 * (r n - l n)" unfolding r_def l_def by (simp add: hom_distribs) finally have "diff (Suc n) \ 3 / 4 * diff n" unfolding diff_def . } note * = this { fix i have "diff i \ (3/4)^i * diff 0" proof (induct i) case (Suc i) from Suc *[of i] show ?case by auto qed auto } then obtain c where *: "\ i. diff i \ (3/4)^i * c" by auto have "\ n. diff n \ eps" proof (cases "c \ 0") case True with *[of 0] eps show ?thesis by (intro exI[of _ 0], auto) next case False hence c: "c > 0" by auto with eps have "inverse c * eps > 0" by auto from exp_tends_to_zero[of "3/4 :: real", OF _ _ this] obtain n where "(3/4) ^ n \ inverse c * eps" by auto from mult_right_mono[OF this, of c] c have "(3/4) ^ n * c \ eps" by (auto simp: field_simps) with *[of n] show ?thesis by (intro exI[of _ n], auto) qed then obtain n where "?r (R n) - ?r (L n) \ eps" unfolding l_def r_def diff_def by blast thus "\n l r. g n = (l, r) \ ?r r - ?r l \ eps" unfolding L_def R_def by (intro exI[of _ n], force) qed qed fun add_1 :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "add_1 (p1,l1,r1) (p2,l2,r2) = ( select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) (\ ((l1,r1,sr1),(l2,r2,sr2)). (l1 + l2, r1 + r2)) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_add p1 p2))" lemma add_1: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z: "z \ add_1 x y" shows "invariant_2 z \ (real_of_2 z = real_of_1 x + real_of_1 y)" proof (cases x) case xt: (fields p1 l1 r1) show ?thesis proof (cases y) case yt: (fields p2 l2 r2) let ?x = "real_of_1 (p1, l1, r1)" let ?y = "real_of_1 (p2, l2, r2)" let ?p = "poly_add p1 p2" note x = x[unfolded xt] note y = y[unfolded yt] from x have ax: "p1 represents ?x" unfolding represents_def by (auto elim!: invariant_1E) from y have ay: "p2 represents ?y" unfolding represents_def by (auto elim!: invariant_1E) let ?bnd = "(\((l1, r1, sr1 :: rat), l2 :: rat, r2 :: rat, sr2 :: rat). (l1 + l2, r1 + r2))" define bnd where "bnd = ?bnd" have "invariant_2 z \ real_of_2 z = ?x + ?y" proof (intro select_correct_factor_int_poly) from represents_add[OF ax ay] show "?p \ 0" "ipoly ?p (?x + ?y) = 0" by auto from z[unfolded xt yt] show sel: "select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) bnd ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_add p1 p2) = z" by (auto simp: bnd_def) have ur1: "unique_root (p1,l1,r1)" "poly_cond2 p1" using x by auto have ur2: "unique_root (p2,l2,r2)" "poly_cond2 p2" using y by auto show "converges_to (\i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))))) (?x + ?y)" by (intro tighten_poly_bounds_binary ur1 ur2; force simp: bnd_def hom_distribs) qed thus ?thesis unfolding xt yt . qed qed declare add_rat_1.simps[simp del] declare add_1.simps[simp del] (* ********************* *) subsubsection\Multiplication\ context begin private fun mult_rat_1_pos :: "rat \ real_alg_1 \ real_alg_2" where "mult_rat_1_pos r1 (p2,l2,r2) = real_alg_2 (cf_pos_poly (poly_mult_rat r1 p2), l2*r1, r2*r1)" private fun mult_1_pos :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "mult_1_pos (p1,l1,r1) (p2,l2,r2) = select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) (\ ((l1,r1,sr1),(l2,r2,sr2)). (l1 * l2, r1 * r2)) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_mult p1 p2)" fun mult_rat_1 :: "rat \ real_alg_1 \ real_alg_2" where "mult_rat_1 x y = (if x < 0 then uminus_2 (mult_rat_1_pos (-x) y) else if x = 0 then Rational 0 else (mult_rat_1_pos x y))" fun mult_1 :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "mult_1 x y = (case (x,y) of ((p1,l1,r1),(p2,l2,r2)) \ if r1 > 0 then if r2 > 0 then mult_1_pos x y else uminus_2 (mult_1_pos x (uminus_1 y)) else if r2 > 0 then uminus_2 (mult_1_pos (uminus_1 x) y) else mult_1_pos (uminus_1 x) (uminus_1 y))" lemma mult_rat_1_pos: fixes r1 :: rat assumes r1: "r1 > 0" and y: "invariant_1 y" defines z: "z \ mult_rat_1_pos r1 y" shows "invariant_2 z \ (real_of_2 z = of_rat r1 * real_of_1 y)" proof - obtain p2 l2 r2 where yt: "y = (p2,l2,r2)" by (cases y, auto) let ?x = "real_of_rat r1" let ?y = "real_of_1 (p2, l2, r2)" let ?p = "poly_mult_rat r1 p2" let ?mp = "cf_pos_poly ?p" note y = y[unfolded yt] note yD = invariant_1D[OF y] from yD r1 have p: "?p \ 0" and r10: "r1 \ 0" by auto hence mp: "?mp \ 0" by simp from yD(1) have rt: "ipoly p2 ?y = 0" and bnd: "of_rat l2 \ ?y" "?y \ of_rat r2" by auto from rt r1 have rt: "ipoly ?mp (?x * ?y) = 0" by (auto simp add: field_simps ipoly_mult_rat[OF r10]) from yD(5) have irr: "irreducible p2" unfolding represents_def using y unfolding root_cond_def split by auto from poly_mult_rat_irreducible[OF this _ r10] yD have irr: "irreducible ?mp" by simp from p have mon: "cf_pos ?mp" by auto obtain l r where lr: "l = l2 * r1" "r = r2 * r1" by force from bnd r1 have bnd: "of_rat l \ ?x * ?y" "?x * ?y \ of_rat r" unfolding lr of_rat_mult by auto with rt have rc: "root_cond (?mp,l,r) (?x * ?y)" unfolding root_cond_def by auto have ur: "unique_root (?mp,l,r)" proof (rule ex1I, rule rc) fix z assume "root_cond (?mp,l,r) z" from this[unfolded root_cond_def split] have bndz: "of_rat l \ z" "z \ of_rat r" and rt: "ipoly ?mp z = 0" by auto have "fst (quotient_of r1) \ 0" using quotient_of_div[of r1] r10 by (cases "quotient_of r1", auto) with rt have rt: "ipoly p2 (z * inverse ?x) = 0" by (auto simp: ipoly_mult_rat[OF r10]) from bndz r1 have "of_rat l2 \ z * inverse ?x" "z * inverse ?x \ of_rat r2" unfolding lr of_rat_mult by (auto simp: field_simps) with rt have "root_cond (p2,l2,r2) (z * inverse ?x)" unfolding root_cond_def by auto also note invariant_1_root_cond[OF y] finally have "?y = z * inverse ?x" by auto thus "z = ?x * ?y" using r1 by auto qed from r1 have sgnr: "sgn r = sgn r2" unfolding lr by (cases "r2 = 0"; cases "r2 < 0"; auto simp: mult_neg_pos mult_less_0_iff) from r1 have sgnl: "sgn l = sgn l2" unfolding lr by (cases "l2 = 0"; cases "l2 < 0"; auto simp: mult_neg_pos mult_less_0_iff) from the_unique_root_eqI[OF ur rc] have xy: "?x * ?y = the_unique_root (?mp,l,r)" by auto from z[unfolded yt, simplified, unfolded Let_def lr[symmetric] split] have z: "z = real_alg_2 (?mp, l, r)" by simp have yp2: "p2 represents ?y" using yD unfolding root_cond_def split represents_def by auto with irr mon have pc: "poly_cond ?mp" by (auto simp: poly_cond_def cf_pos_def) have rc: "invariant_1 (?mp, l, r)" unfolding z using yD(2) pc ur by (auto simp add: invariant_1_def ur mp sgnr sgnl) show ?thesis unfolding z using real_alg_2[OF rc] unfolding yt xy unfolding z by simp qed lemma mult_1_pos: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z: "z \ mult_1_pos x y" assumes pos: "real_of_1 x > 0" "real_of_1 y > 0" shows "invariant_2 z \ (real_of_2 z = real_of_1 x * real_of_1 y)" proof - obtain p1 l1 r1 where xt: "x = (p1,l1,r1)" by (cases x, auto) obtain p2 l2 r2 where yt: "y = (p2,l2,r2)" by (cases y, auto) let ?x = "real_of_1 (p1, l1, r1)" let ?y = "real_of_1 (p2, l2, r2)" let ?r = "real_of_rat" let ?p = "poly_mult p1 p2" note x = x[unfolded xt] note y = y[unfolded yt] from x y have basic: "unique_root (p1, l1, r1)" "poly_cond2 p1" "unique_root (p2, l2, r2)" "poly_cond2 p2" by auto from basic have irr1: "irreducible p1" and irr2: "irreducible p2" by auto from x have ax: "p1 represents ?x" unfolding represents_def by (auto elim!:invariant_1E) from y have ay: "p2 represents ?y" unfolding represents_def by (auto elim!:invariant_1E) from ax ay pos[unfolded xt yt] have axy: "?p represents (?x * ?y)" by (intro represents_mult represents_irr_non_0[OF irr2], auto) from representsD[OF this] have p: "?p \ 0" and rt: "ipoly ?p (?x * ?y) = 0" . from x pos(1)[unfolded xt] have "?r r1 > 0" unfolding split by auto hence "sgn r1 = 1" unfolding sgn_rat_def by (auto split: if_splits) with x have "sgn l1 = 1" by auto hence l1_pos: "l1 > 0" unfolding sgn_rat_def by (cases "l1 = 0"; cases "l1 < 0"; auto) from y pos(2)[unfolded yt] have "?r r2 > 0" unfolding split by auto hence "sgn r2 = 1" unfolding sgn_rat_def by (auto split: if_splits) with y have "sgn l2 = 1" by auto hence l2_pos: "l2 > 0" unfolding sgn_rat_def by (cases "l2 = 0"; cases "l2 < 0"; auto) let ?bnd = "(\((l1, r1, sr1 :: rat), l2 :: rat, r2 :: rat, sr2 :: rat). (l1 * l2, r1 * r2))" define bnd where "bnd = ?bnd" obtain z' where sel: "select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) bnd ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) ?p = z'" by auto have main: "invariant_2 z' \ real_of_2 z' = ?x * ?y" proof (rule select_correct_factor_int_poly[OF _ sel rt p]) { fix l1 r1 l2 r2 l1' r1' l2' r2' l l' r r' :: rat let ?m1 = "(l1+r1)/2" let ?m2 = "(l2+r2)/2" define d1 where "d1 = r1 - l1" define d2 where "d2 = r2 - l2" let ?M1 = "l1 + d1/2" let ?M2 = "l2 + d2/2" assume le: "l1 > 0" "l2 > 0" "l1 \ r1" "l2 \ r2" and id: "(l, r) = (l1 * l2, r1 * r2)" "(l', r') = (l1' * l2', r1' * r2')" and mem: "(l1', r1') \ {(l1, ?m1), (?m1, r1)}" "(l2', r2') \ {(l2, ?m2), (?m2, r2)}" hence id: "l = l1 * l2" "r = (l1 + d1) * (l2 + d2)" "l' = l1' * l2'" "r' = r1' * r2'" "r1 = l1 + d1" "r2 = l2 + d2" and id': "?m1 = ?M1" "?m2 = ?M2" unfolding d1_def d2_def by (auto simp: field_simps) define l1d1 where "l1d1 = l1 + d1" from le have ge0: "d1 \ 0" "d2 \ 0" "l1 \ 0" "l2 \ 0" unfolding d1_def d2_def by auto have "4 * (r' - l') \ 3 * (r - l)" proof (cases "l1' = l1 \ r1' = ?M1 \ l2' = l2 \ r2' = ?M2") case True hence id2: "l1' = l1" "r1' = ?M1" "l2' = l2" "r2' = ?M2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 1 = this show ?thesis proof (cases "l1' = l1 \ r1' = ?M1 \ l2' = ?M2 \ r2' = r2") case True hence id2: "l1' = l1" "r1' = ?M1" "l2' = ?M2" "r2' = r2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 2 = this show ?thesis proof (cases "l1' = ?M1 \ r1' = r1 \ l2' = l2 \ r2' = ?M2") case True hence id2: "l1' = ?M1" "r1' = r1" "l2' = l2" "r2' = ?M2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 3 = this from 1 2 3 mem have id2: "l1' = ?M1" "r1' = r1" "l2' = ?M2" "r2' = r2" unfolding id' by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp qed qed qed hence "r' - l' \ 3 / 4 * (r - l)" by simp } note decr = this show "converges_to (\i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))))) (?x * ?y)" proof (intro tighten_poly_bounds_binary[where f = "(*)" and I = "\ l. l > 0"] basic l1_pos l2_pos, goal_cases) case (1 L1 R1 L2 R2 L R) hence "L = L1 * L2" "R = R1 * R2" unfolding bnd_def by auto hence id: "?r L = ?r L1 * ?r L2" "?r R = ?r R1 * ?r R2" by (auto simp: hom_distribs) from 1(3-4) have le: "?r L1 \ ?x" "?x \ ?r R1" "?r L2 \ ?y" "?y \ ?r R2" unfolding root_cond_def by auto from 1(1-2) have lt: "0 < ?r L1" "0 < ?r L2" by auto from mult_mono[OF le(1,3), folded id] lt le have L: "?r L \ ?x * ?y" by linarith have R: "?x * ?y \ ?r R" by (rule mult_mono[OF le(2,4), folded id], insert lt le, linarith+) show ?case using L R by blast next case (2 l1 r1 l2 r2 l1' r1' l2' r2' l l' r r') from 2(5-6) have lr: "l = l1 * l2" "r = r1 * r2" "l' = l1' * l2'" "r' = r1' * r2'" unfolding bnd_def by auto from 2(1-4) have le: "0 < l1" "0 < l2" "l1 \ r1" "l2 \ r2" by auto from 2(7-8) le have le': "l1 \ l1'" "r1' \ r1" "l2 \ l2'" "r2' \ r2" "0 < r2'" "0 < r2" by auto from mult_mono[OF le'(1,3), folded lr] le le' have l: "l \ l'" by auto have r: "r' \ r" by (rule mult_mono[OF le'(2,4), folded lr], insert le le', linarith+) have "r' - l' \ 3 / 4 * (r - l)" by (rule decr[OF _ _ _ _ _ _ 2(7-8)], insert le le' lr, auto) thus ?case using l r by blast qed auto qed have z': "z' = z" unfolding z[unfolded xt yt, simplified, unfolded bnd_def[symmetric] sel] by auto from main[unfolded this] show ?thesis unfolding xt yt by simp qed lemma mult_1: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z[simp]: "z \ mult_1 x y" shows "invariant_2 z \ (real_of_2 z = real_of_1 x * real_of_1 y)" proof - obtain p1 l1 r1 where xt[simp]: "x = (p1,l1,r1)" by (cases x) obtain p2 l2 r2 where yt[simp]: "y = (p2,l2,r2)" by (cases y) let ?xt = "(p1, l1, r1)" let ?yt = "(p2, l2, r2)" let ?x = "real_of_1 ?xt" let ?y = "real_of_1 ?yt" let ?mxt = "uminus_1 ?xt" let ?myt = "uminus_1 ?yt" let ?mx = "real_of_1 ?mxt" let ?my = "real_of_1 ?myt" let ?r = "real_of_rat" from invariant_1_2_of_rat[OF x, of 0] have x0: "?x < 0 \ ?x > 0" by auto from invariant_1_2_of_rat[OF y, of 0] have y0: "?y < 0 \ ?y > 0" by auto from uminus_1_2[OF x] have mx: "invariant_1_2 ?mxt" and [simp]: "?mx = - ?x" by auto from uminus_1_2[OF y] have my: "invariant_1_2 ?myt" and [simp]: "?my = - ?y" by auto have id: "r1 > 0 \ ?x > 0" "r1 < 0 \ ?x < 0" "r2 > 0 \ ?y > 0" "r2 < 0 \ ?y < 0" using x y by auto show ?thesis proof (cases "?x > 0") case x0: True show ?thesis proof (cases "?y > 0") case y0: True with x y x0 mult_1_pos[OF x y] show ?thesis by auto next case False with y0 have y0: "?y < 0" by auto with x0 have z: "z = uminus_2 (mult_1_pos ?xt ?myt)" unfolding z xt yt mult_1.simps split id by simp from x0 y0 mult_1_pos[OF x my] uminus_2[of "mult_1_pos ?xt ?myt"] show ?thesis unfolding z by simp qed next case False with x0 have x0: "?x0 < 0" by simp show ?thesis proof (cases "?y > 0") case y0: True with x0 x y id have z: "z = uminus_2 (mult_1_pos ?mxt ?yt)" by simp from x0 y0 mult_1_pos[OF mx y] uminus_2[of "mult_1_pos ?mxt ?yt"] show ?thesis unfolding z by auto next case False with y0 have y0: "?y < 0" by simp with x0 x y have z: "z = mult_1_pos ?mxt ?myt" by auto with x0 y0 x y mult_1_pos[OF mx my] show ?thesis unfolding z by auto qed qed qed lemma mult_rat_1: fixes x assumes y: "invariant_1 y" defines z: "z \ mult_rat_1 x y" shows "invariant_2 z \ (real_of_2 z = of_rat x * real_of_1 y)" proof (cases y) case yt: (fields p2 l2 r2) let ?yt = "(p2, l2, r2)" let ?x = "real_of_rat x" let ?y = "real_of_1 ?yt" let ?myt = "mult_rat_1_pos (- x) ?yt" note y = y[unfolded yt] note z = z[unfolded yt] show ?thesis proof(cases x "0::rat" rule:linorder_cases) case x: greater with z have z: "z = mult_rat_1_pos x ?yt" by simp from mult_rat_1_pos[OF x y] show ?thesis unfolding yt z by auto next case less then have x: "- x > 0" by auto hence z: "z = uminus_2 ?myt" unfolding z by simp from mult_rat_1_pos[OF x y] have rc: "invariant_2 ?myt" and rr: "real_of_2 ?myt = - ?x * ?y" by (auto simp: hom_distribs) from uminus_2[OF rc] rr show ?thesis unfolding z[symmetric] unfolding yt[symmetric] by simp qed (auto simp: z) qed end declare mult_1.simps[simp del] declare mult_rat_1.simps[simp del] (* ********************* *) subsubsection\Root\ definition ipoly_root_delta :: "int poly \ real" where "ipoly_root_delta p = Min (insert 1 { abs (x - y) | x y. ipoly p x = 0 \ ipoly p y = 0 \ x \ y}) / 4" lemma ipoly_root_delta: assumes "p \ 0" shows "ipoly_root_delta p > 0" "2 \ card (Collect (root_cond (p, l, r))) \ ipoly_root_delta p \ real_of_rat (r - l) / 4" proof - let ?z = "0 :: real" let ?R = "{x. ipoly p x = ?z}" let ?set = "{ abs (x - y) | x y. ipoly p x = ?z \ ipoly p y = 0 \ x \ y}" define S where "S = insert 1 ?set" from finite_ipoly_roots[OF assms] have finR: "finite ?R" and fin: "finite (?R \ ?R)" by auto have "finite ?set" by (rule finite_subset[OF _ finite_imageI[OF fin, of "\ (x,y). abs (x - y)"]], force) hence fin: "finite S" and ne: "S \ {}" and pos: "\ x. x \ S \ x > 0" unfolding S_def by auto have delta: "ipoly_root_delta p = Min S / 4" unfolding ipoly_root_delta_def S_def .. have pos: "Min S > 0" using fin ne pos by auto show "ipoly_root_delta p > 0" unfolding delta using pos by auto let ?S = "Collect (root_cond (p, l, r))" assume "2 \ card ?S" hence 2: "Suc (Suc 0) \ card ?S" by simp from 2[unfolded card_le_Suc_iff[of _ ?S]] obtain x T where ST: "?S = insert x T" and xT: "x \ T" and 1: "Suc 0 \ card T" by auto from 1[unfolded card_le_Suc_iff[of _ T]] obtain y where yT: "y \ T" by auto from ST xT yT have x: "x \ ?S" and y: "y \ ?S" and xy: "x \ y" by auto hence "abs (x - y) \ S" unfolding S_def root_cond_def[abs_def] by auto with fin have "Min S \ abs (x - y)" by auto with pos have le: "Min S / 2 \ abs (x - y) / 2" by auto from x y have "abs (x - y) \ of_rat r - of_rat l" unfolding root_cond_def[abs_def] by auto also have "\ = of_rat (r - l)" by (auto simp: of_rat_diff) finally have "abs (x - y) / 2 \ of_rat (r - l) / 2" by auto with le show "ipoly_root_delta p \ real_of_rat (r - l) / 4" unfolding delta by auto qed lemma sgn_less_eq_1_rat: fixes a b :: rat shows "sgn a = 1 \ a \ b \ sgn b = 1" by (metis (no_types, opaque_lifting) not_less one_neq_neg_one one_neq_zero order_trans sgn_rat_def) lemma sgn_less_eq_1_real: fixes a b :: real shows "sgn a = 1 \ a \ b \ sgn b = 1" by (metis (no_types, opaque_lifting) not_less one_neq_neg_one one_neq_zero order_trans sgn_real_def) definition compare_1_rat :: "real_alg_1 \ rat \ order" where "compare_1_rat rai = (let p = poly_real_alg_1 rai in if degree p = 1 then let x = Rat.Fract (- coeff p 0) (coeff p 1) in (\ y. compare y x) else (\ y. compare_rat_1 y rai))" lemma compare_real_of_rat: "compare (real_of_rat x) (of_rat y) = compare x y" unfolding compare_rat_def compare_real_def comparator_of_def of_rat_less by auto lemma compare_1_rat: assumes rc: "invariant_1 y" shows "compare_1_rat y x = compare (of_rat x) (real_of_1 y)" proof (cases "degree (poly_real_alg_1 y)" "Suc 0" rule: linorder_cases) case less with invariant_1_degree_0[OF rc] show ?thesis by auto next case deg: greater with rc have rc: "invariant_1_2 y" by auto from deg compare_rat_1[OF rc, of x] show ?thesis unfolding compare_1_rat_def by auto next case deg: equal obtain p l r where y: "y = (p,l,r)" by (cases y) note rc = invariant_1D[OF rc[unfolded y]] from deg have p: "degree p = Suc 0" and id: "compare_1_rat y x = compare x (Rat.Fract (- coeff p 0) (coeff p 1))" unfolding compare_1_rat_def by (auto simp: Let_def y) from rc(1)[unfolded split] have "ipoly p (real_of_1 y) = 0" unfolding y by auto with degree_1_ipoly[OF p, of "real_of_1 y"] have id': "real_of_1 y = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1))" by simp show ?thesis unfolding id id' compare_real_of_rat .. qed context fixes n :: nat begin private definition initial_lower_bound :: "rat \ rat" where "initial_lower_bound l = (if l \ 1 then l else of_int (root_rat_floor n l))" private definition initial_upper_bound :: "rat \ rat" where "initial_upper_bound r = (of_int (root_rat_ceiling n r))" context fixes cmpx :: "rat \ order" begin fun tighten_bound_root :: "rat \ rat \ rat \ rat" where "tighten_bound_root (l',r') = (let m' = (l' + r') / 2; m = m' ^ n in case cmpx m of Eq \ (m',m') | Lt \ (m',r') | Gt \ (l',m'))" lemma tighten_bound_root: assumes sgn: "sgn il = 1" "real_of_1 x \ 0" and il: "real_of_rat il \ root n (real_of_1 x)" and ir: "root n (real_of_1 x) \ real_of_rat ir" and rai: "invariant_1 x" and cmpx: "cmpx = compare_1_rat x" and n: "n \ 0" shows "converges_to (\ i. (tighten_bound_root ^^ i) (il, ir)) (root n (real_of_1 x))" (is "converges_to ?f ?x") unfolding converges_to_def proof (intro conjI impI allI) { fix x :: real have "x \ 0 \ (root n x) ^ n = x" using n by simp } note root_exp_cancel = this { fix x :: real have "x \ 0 \ root n (x ^ n) = x" using n using real_root_pos_unique by blast } note root_exp_cancel' = this from il ir have "real_of_rat il \ of_rat ir" by auto hence ir_il: "il \ ir" by (auto simp: of_rat_less_eq) from n have n': "n > 0" by auto { fix i have "in_interval (?f i) ?x \ sub_interval (?f i) (il,ir) \ (i \ 0 \ sub_interval (?f i) (?f (i - 1))) \ snd (?f i) - fst (?f i) \ (ir - il) / 2^i" proof (induct i) case 0 show ?case using il ir by auto next case (Suc i) obtain l' r' where id: "(tighten_bound_root ^^ i) (il, ir) = (l',r')" by (cases "(tighten_bound_root ^^ i) (il, ir)", auto) let ?m' = "(l' + r') / 2" let ?m = "?m' ^ n" define m where "m = ?m" note IH = Suc[unfolded id split snd_conv fst_conv] from IH have "sub_interval (l', r') (il, ir)" by auto hence ill': "il \ l'" "r' \ ir" by auto with sgn have l'0: "l' > 0" using sgn_1_pos sgn_less_eq_1_rat by blast from IH have lr'x: "in_interval (l', r') ?x" by auto hence lr'': "real_of_rat l' \ of_rat r'" by auto hence lr': "l' \ r'" unfolding of_rat_less_eq . with l'0 have r'0: "r' > 0" by auto note compare = compare_1_rat[OF rai, of ?m, folded cmpx] from IH have *: "r' - l' \ (ir - il) / 2 ^ i" by auto have "r' - (l' + r') / 2 = (r' - l') / 2" by (simp add: field_simps) also have "\ \ (ir - il) / 2 ^ i / 2" using * by (rule divide_right_mono, auto) finally have size: "r' - (l' + r') / 2 \ (ir - il) / (2 * 2 ^ i)" by simp also have "r' - (l' + r') / 2 = (l' + r') / 2 - l'" by auto finally have size': "(l' + r') / 2 - l' \ (ir - il) / (2 * 2 ^ i)" by simp have "root n (real_of_rat ?m) = root n ((real_of_rat ?m') ^ n)" by (simp add: hom_distribs) also have "\ = real_of_rat ?m'" by (rule root_exp_cancel', insert l'0 lr', auto) finally have root: "root n (of_rat ?m) = of_rat ?m'" . show ?case proof (cases "cmpx ?m") case Eq from compare[unfolded Eq] have "real_of_1 x = of_rat ?m" unfolding compare_real_def comparator_of_def by (auto split: if_splits) from arg_cong[OF this, of "root n"] have "?x = root n (of_rat ?m)" . also have "\ = root n (real_of_rat ?m') ^ n" using n real_root_power by (auto simp: hom_distribs) also have "\ = of_rat ?m'" by (rule root_exp_cancel, insert IH sgn(2) l'0 r'0, auto) finally have x: "?x = of_rat ?m'" . show ?thesis using x id Eq lr' ill' ir_il by (auto simp: Let_def) next case Lt from compare[unfolded Lt] have lt: "of_rat ?m \ real_of_1 x" unfolding compare_real_def comparator_of_def by (auto split: if_splits) have id'': "?f (Suc i) = (?m',r')" "?f (Suc i - 1) = (l',r')" using Lt id by (auto simp add: Let_def) from real_root_le_mono[OF n' lt] have "of_rat ?m' \ ?x" unfolding root by simp with lr'x lr'' have ineq': "real_of_rat l' + real_of_rat r' \ ?x * 2" by (auto simp: hom_distribs) show ?thesis unfolding id'' by (auto simp: Let_def hom_distribs, insert size ineq' lr' ill' lr'x ir_il, auto) next case Gt from compare[unfolded Gt] have lt: "of_rat ?m \ real_of_1 x" unfolding compare_real_def comparator_of_def by (auto split: if_splits) have id'': "?f (Suc i) = (l',?m')" "?f (Suc i - 1) = (l',r')" using Gt id by (auto simp add: Let_def) from real_root_le_mono[OF n' lt] have "?x \ of_rat ?m'" unfolding root by simp with lr'x lr'' have ineq': "?x * 2 \ real_of_rat l' + real_of_rat r'" by (auto simp: hom_distribs) show ?thesis unfolding id'' by (auto simp: Let_def hom_distribs, insert size' ineq' lr' ill' lr'x ir_il, auto) qed qed } note main = this fix i from main[of i] show "in_interval (?f i) ?x" by auto from main[of "Suc i"] show "sub_interval (?f (Suc i)) (?f i)" by auto fix eps :: real assume eps: "0 < eps" define c where "c = eps / (max (real_of_rat (ir - il)) 1)" have c0: "c > 0" using eps unfolding c_def by auto from exp_tends_to_zero[OF _ _ this, of "1/2"] obtain i where c: "(1/2)^i \ c" by auto obtain l' r' where fi: "?f i = (l',r')" by force from main[of i, unfolded fi] have le: "r' - l' \ (ir - il) / 2 ^ i" by auto have iril: "real_of_rat (ir - il) \ 0" using ir_il by (auto simp: of_rat_less_eq) show "\n la ra. ?f n = (la, ra) \ real_of_rat ra - real_of_rat la \ eps" proof (intro conjI exI, rule fi) have "real_of_rat r' - of_rat l' = real_of_rat (r' - l')" by (auto simp: hom_distribs) also have "\ \ real_of_rat ((ir - il) / 2 ^ i)" using le unfolding of_rat_less_eq . also have "\ = (real_of_rat (ir - il)) * ((1/2) ^ i)" by (simp add: field_simps hom_distribs) also have "\ \ (real_of_rat (ir - il)) * c" by (rule mult_left_mono[OF c iril]) also have "\ \ eps" proof (cases "real_of_rat (ir - il) \ 1") case True hence "c = eps" unfolding c_def by (auto simp: hom_distribs) thus ?thesis using eps True by auto next case False hence "max (real_of_rat (ir - il)) 1 = real_of_rat (ir - il)" "real_of_rat (ir - il) \ 0" by (auto simp: hom_distribs) hence "(real_of_rat (ir - il)) * c = eps" unfolding c_def by auto thus ?thesis by simp qed finally show "real_of_rat r' - of_rat l' \ eps" . qed qed end private fun root_pos_1 :: "real_alg_1 \ real_alg_2" where "root_pos_1 (p,l,r) = ( (select_correct_factor_int_poly (tighten_bound_root (compare_1_rat (p,l,r))) (\ x. x) (initial_lower_bound l, initial_upper_bound r) (poly_nth_root n p)))" fun root_1 :: "real_alg_1 \ real_alg_2" where "root_1 (p,l,r) = ( if n = 0 \ r = 0 then Rational 0 else if r > 0 then root_pos_1 (p,l,r) else uminus_2 (root_pos_1 (uminus_1 (p,l,r))))" context assumes n: "n \ 0" begin lemma initial_upper_bound: assumes x: "x > 0" and xr: "x \ of_rat r" shows "sgn (initial_upper_bound r) = 1" "root n x \ of_rat (initial_upper_bound r)" proof - have n: "n > 0" using n by auto note d = initial_upper_bound_def let ?r = "initial_upper_bound r" from x xr have r0: "r > 0" by (meson not_less of_rat_le_0_iff order_trans) hence "of_rat r > (0 :: real)" by auto hence "root n (of_rat r) > 0" using n by simp hence "1 \ ceiling (root n (of_rat r))" by auto hence "(1 :: rat) \ of_int (ceiling (root n (of_rat r)))" by linarith also have "\ = ?r" unfolding d by simp finally show "sgn ?r = 1" unfolding sgn_rat_def by auto have "root n x \ root n (of_rat r)" unfolding real_root_le_iff[OF n] by (rule xr) also have "\ \ of_rat ?r" unfolding d by simp finally show "root n x \ of_rat ?r" . qed lemma initial_lower_bound: assumes l: "l > 0" and lx: "of_rat l \ x" shows "sgn (initial_lower_bound l) = 1" "of_rat (initial_lower_bound l) \ root n x" proof - have n: "n > 0" using n by auto note d = initial_lower_bound_def let ?l = "initial_lower_bound l" from l lx have x0: "x > 0" by (meson not_less of_rat_le_0_iff order_trans) have "sgn ?l = 1 \ of_rat ?l \ root n x" proof (cases "l \ 1") case True hence ll: "?l = l" and l0: "of_rat l \ (0 :: real)" and l1: "of_rat l \ (1 :: real)" using l unfolding True d by auto have sgn: "sgn ?l = 1" using l unfolding ll by auto have "of_rat ?l = of_rat l" unfolding ll by simp also have "of_rat l \ root n (of_rat l)" using real_root_increasing[OF _ _ l0 l1, of 1 n] n by (cases "n = 1", auto) also have "\ \ root n x" using lx unfolding real_root_le_iff[OF n] . finally show ?thesis using sgn by auto next case False hence l: "(1 :: real) \ of_rat l" and ll: "?l = of_int (floor (root n (of_rat l)))" unfolding d by auto hence "root n 1 \ root n (of_rat l)" unfolding real_root_le_iff[OF n] by auto hence "1 \ root n (of_rat l)" using n by auto from floor_mono[OF this] have "1 \ ?l" using one_le_floor unfolding ll by fastforce hence sgn: "sgn ?l = 1" by simp have "of_rat ?l \ root n (of_rat l)" unfolding ll by simp also have "\ \ root n x" using lx unfolding real_root_le_iff[OF n] . finally have "of_rat ?l \ root n x" . with sgn show ?thesis by auto qed thus "sgn ?l = 1" "of_rat ?l \ root n x" by auto qed lemma root_pos_1: assumes x: "invariant_1 x" and pos: "rai_ub x > 0" defines y: "y \ root_pos_1 x" shows "invariant_2 y \ real_of_2 y = root n (real_of_1 x)" proof (cases x) case (fields p l r) let ?l = "initial_lower_bound l" let ?r = "initial_upper_bound r" from x fields have rai: "invariant_1 (p,l,r)" by auto note * = invariant_1D[OF this] let ?x = "the_unique_root (p,l,r)" from pos[unfolded fields] * have sgnl: "sgn l = 1" by auto from sgnl have l0: "l > 0" by (unfold sgn_1_pos) hence ll0: "real_of_rat l > 0" by auto from * have lx: "of_rat l \ ?x" by auto with ll0 have x0: "?x > 0" by linarith note il = initial_lower_bound[OF l0 lx] from * have "?x \ of_rat r" by auto note iu = initial_upper_bound[OF x0 this] let ?p = "poly_nth_root n p" from x0 have id: "root n ?x ^ n = ?x" using n real_root_pow_pos by blast have rc: "root_cond (?p, ?l, ?r) (root n ?x)" using il iu * by (intro root_condI, auto simp: ipoly_nth_root id) hence root: "ipoly ?p (root n (real_of_1 x)) = 0" unfolding root_cond_def fields by auto from * have "p \ 0" by auto hence p': "?p \ 0" using poly_nth_root_0[of n p] n by auto have tbr: "0 \ real_of_1 x" "real_of_rat (initial_lower_bound l) \ root n (real_of_1 x)" "root n (real_of_1 x) \ real_of_rat (initial_upper_bound r)" using x0 il(2) iu(2) fields by auto from select_correct_factor_int_poly[OF tighten_bound_root[OF il(1)[folded fields] tbr x refl n] refl root p'] show ?thesis by (simp add: y fields) qed end lemma root_1: assumes x: "invariant_1 x" defines y: "y \ root_1 x" shows "invariant_2 y \ (real_of_2 y = root n (real_of_1 x))" proof (cases "n = 0 \ rai_ub x = 0") case True with x have "n = 0 \ real_of_1 x = 0" by (cases x, auto) then have "root n (real_of_1 x) = 0" by auto then show ?thesis unfolding y root_1.simps using x by (cases x, auto) next case False with x have n: "n \ 0" and x0: "real_of_1 x \ 0" by (simp, cases x, auto) note rt = root_pos_1 show ?thesis proof (cases "rai_ub x" "0::rat" rule:linorder_cases) case greater with rt[OF n x this] n show ?thesis by (unfold y, cases x, simp) next case less let ?um = "uminus_1" let ?rt = "root_pos_1" from n less y x0 have y: "y = uminus_2 (?rt (?um x))" by (cases x, auto) from uminus_1[OF x] have umx: "invariant_1 (?um x)" and umx2: "real_of_1 (?um x) = - real_of_1 x" by auto with x less have "0 < rai_ub (uminus_1 x)" by (cases x, auto simp: uminus_1.simps Let_def) from rt[OF n umx this] umx2 have rumx: "invariant_2 (?rt (?um x))" and rumx2: "real_of_2 (?rt (?um x)) = root n (- real_of_1 x)" by auto from uminus_2[OF rumx] rumx2 y real_root_minus show ?thesis by auto next case equal with x0 x show ?thesis by (cases x, auto) qed qed end declare root_1.simps[simp del] (* ********************** *) subsubsection \Embedding of Rational Numbers\ definition of_rat_1 :: "rat \ real_alg_1" where "of_rat_1 x \ (poly_rat x,x,x)" lemma of_rat_1: shows "invariant_1 (of_rat_1 x)" and "real_of_1 (of_rat_1 x) = of_rat x" unfolding of_rat_1_def by (atomize(full), intro invariant_1_realI unique_rootI poly_condI, auto ) fun info_2 :: "real_alg_2 \ rat + int poly \ nat" where "info_2 (Rational x) = Inl x" | "info_2 (Irrational n (p,l,r)) = Inr (p,n)" lemma info_2_card: assumes rc: "invariant_2 x" shows "info_2 x = Inr (p,n) \ poly_cond p \ ipoly p (real_of_2 x) = 0 \ degree p \ 2 \ card (roots_below p (real_of_2 x)) = n" "info_2 x = Inl y \ real_of_2 x = of_rat y" proof (atomize(full), goal_cases) case 1 show ?case proof (cases x) case (Irrational m rai) then obtain q l r where x: "x = Irrational m (q,l,r)" by (cases rai, auto) show ?thesis proof (cases "q = p \ m = n") case False thus ?thesis using x by auto next case True with x have x: "x = Irrational n (p,l,r)" by auto from rc[unfolded x, simplified] have inv: "invariant_1_2 (p,l,r)" and n: "card (roots_below p (real_of_2 x)) = n" and 1: "degree p \ 1" by (auto simp: x) from inv have "degree p \ 0" unfolding irreducible_def by auto with 1 have "degree p \ 2" by linarith thus ?thesis unfolding n using inv x by (auto elim!: invariant_1E) qed qed auto qed lemma real_of_2_Irrational: "invariant_2 (Irrational n rai) \ real_of_2 (Irrational n rai) \ of_rat x" proof assume "invariant_2 (Irrational n rai)" and rat: "real_of_2 (Irrational n rai) = real_of_rat x" hence "real_of_1 rai \ \" "invariant_1_2 rai" by auto from invariant_1_2_of_rat[OF this(2)] rat show False by auto qed lemma info_2: assumes ix: "invariant_2 x" and iy: "invariant_2 y" shows "info_2 x = info_2 y \ real_of_2 x = real_of_2 y" proof (cases x) case x: (Irrational n1 rai1) note ix = ix[unfolded x] show ?thesis proof (cases y) case (Rational y) with real_of_2_Irrational[OF ix, of y] show ?thesis unfolding x by (cases rai1, auto) next case y: (Irrational n2 rai2) obtain p1 l1 r1 where rai1: "rai1 = (p1,l1,r1)" by (cases rai1) obtain p2 l2 r2 where rai2: "rai2 = (p2,l2,r2)" by (cases rai2) let ?rx = "the_unique_root (p1,l1,r1)" let ?ry = "the_unique_root (p2,l2,r2)" have id: "(info_2 x = info_2 y) = (p1 = p2 \ n1 = n2)" "(real_of_2 x = real_of_2 y) = (?rx = ?ry)" unfolding x y rai1 rai2 by auto from ix[unfolded x rai1] have ix: "invariant_1 (p1, l1, r1)" and deg1: "degree p1 > 1" and n1: "n1 = card (roots_below p1 ?rx)" by auto note Ix = invariant_1D[OF ix] from deg1 have p1_0: "p1 \ 0" by auto from iy[unfolded y rai2] have iy: "invariant_1 (p2, l2, r2)" and "degree p2 > 1" and n2: "n2 = card (roots_below p2 ?ry)" by auto note Iy = invariant_1D[OF iy] show ?thesis unfolding id proof assume eq: "?rx = ?ry" from Ix have algx: "p1 represents ?rx \ irreducible p1 \ lead_coeff p1 > 0" unfolding represents_def by auto from iy have algy: "p2 represents ?rx \ irreducible p2 \ lead_coeff p2 > 0" unfolding represents_def eq by (auto elim!: invariant_1E) from algx have "algebraic ?rx" unfolding algebraic_altdef_ipoly by auto note unique = algebraic_imp_represents_unique[OF this] with algx algy have id: "p2 = p1" by auto from eq id n1 n2 show "p1 = p2 \ n1 = n2" by auto next assume "p1 = p2 \ n1 = n2" hence id: "p1 = p2" "n1 = n2" by auto hence card: "card (roots_below p1 ?rx) = card (roots_below p1 ?ry)" unfolding n1 n2 by auto show "?rx = ?ry" proof (cases ?rx ?ry rule: linorder_cases) case less have "roots_below p1 ?rx = roots_below p1 ?ry" proof (intro card_subset_eq finite_subset[OF _ ipoly_roots_finite] card) from less show "roots_below p1 ?rx \ roots_below p1 ?ry" by auto qed (insert p1_0, auto) then show ?thesis using id less unique_rootD(3)[OF Iy(4)] by (auto simp: less_eq_real_def) next case equal then show ?thesis by (simp add: id) next case greater have "roots_below p1 ?ry = roots_below p1 ?rx" proof (intro card_subset_eq card[symmetric] finite_subset[OF _ ipoly_roots_finite[OF p1_0]]) from greater show "roots_below p1 ?ry \ roots_below p1 ?rx" by auto qed auto hence "roots_below p2 ?ry = roots_below p2 ?rx" unfolding id by auto thus ?thesis using id greater unique_rootD(3)[OF Ix(4)] by (auto simp: less_eq_real_def) qed qed qed next case x: (Rational x) show ?thesis proof (cases y) case (Rational y) thus ?thesis using x by auto next case y: (Irrational n rai) with real_of_2_Irrational[OF iy[unfolded y], of x] show ?thesis unfolding x by (cases rai, auto) qed qed lemma info_2_unique: "invariant_2 x \ invariant_2 y \ real_of_2 x = real_of_2 y \ info_2 x = info_2 y" using info_2 by blast lemma info_2_inj: "invariant_2 x \ invariant_2 y \ info_2 x = info_2 y \ real_of_2 x = real_of_2 y" using info_2 by blast context fixes cr1 cr2 :: "rat \ rat \ nat" begin partial_function (tailrec) compare_1 :: "int poly \ int poly \ rat \ rat \ rat \ rat \ rat \ rat \ order" where [code]: "compare_1 p1 p2 l1 r1 sr1 l2 r2 sr2 = (if r1 < l2 then Lt else if r2 < l1 then Gt else let (l1',r1',sr1') = tighten_poly_bounds p1 l1 r1 sr1; (l2',r2',sr2') = tighten_poly_bounds p2 l2 r2 sr2 in compare_1 p1 p2 l1' r1' sr1' l2' r2' sr2') " lemma compare_1: assumes ur1: "unique_root (p1,l1,r1)" and ur2: "unique_root (p2,l2,r2)" and pc: "poly_cond2 p1" "poly_cond2 p2" and diff: "the_unique_root (p1,l1,r1) \ the_unique_root (p2,l2,r2)" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" shows "compare_1 p1 p2 l1 r1 sr1 l2 r2 sr2 = compare (the_unique_root (p1,l1,r1)) (the_unique_root (p2,l2,r2))" proof - let ?r = real_of_rat { fix d x y assume d: "d = (r1 - l1) + (r2 - l2)" and xy: "x = the_unique_root (p1,l1,r1)" "y = the_unique_root (p2,l2,r2)" define delta where "delta = abs (x - y) / 4" have delta: "delta > 0" and diff: "x \ y" unfolding delta_def using diff xy by auto let ?rel' = "{(x, y). 0 \ y \ delta_gt delta x y}" let ?rel = "inv_image ?rel' ?r" have SN: "SN ?rel" by (rule SN_inv_image[OF delta_gt_SN[OF delta]]) from d ur1 ur2 have ?thesis unfolding xy[symmetric] using xy sr proof (induct d arbitrary: l1 r1 l2 r2 sr1 sr2 rule: SN_induct[OF SN]) case (1 d l1 r1 l2 r2) note IH = 1(1) note d = 1(2) note ur = 1(3-4) note xy = 1(5-6) note sr = 1(7-8) note simps = compare_1.simps[of p1 p2 l1 r1 sr1 l2 r2 sr2] note urx = unique_rootD[OF ur(1), folded xy] note ury = unique_rootD[OF ur(2), folded xy] show ?case (is "?l = _") proof (cases "r1 < l2") case True hence l: "?l = Lt" and lt: "?r r1 < ?r l2" unfolding simps of_rat_less by auto show ?thesis unfolding l using lt True urx(2) ury(1) by (auto simp: compare_real_def comparator_of_def) next case False note le = this show ?thesis proof (cases "r2 < l1") case True with le have l: "?l = Gt" and lt: "?r r2 < ?r l1" unfolding simps of_rat_less by auto show ?thesis unfolding l using lt True ury(2) urx(1) by (auto simp: compare_real_def comparator_of_def) next case False obtain l1' r1' sr1' where tb1: "tighten_poly_bounds p1 l1 r1 sr1 = (l1',r1',sr1')" by (cases rule: prod_cases3, auto) obtain l2' r2' sr2' where tb2: "tighten_poly_bounds p2 l2 r2 sr2 = (l2',r2',sr2')" by (cases rule: prod_cases3, auto) from False le tb1 tb2 have l: "?l = compare_1 p1 p2 l1' r1' sr1' l2' r2' sr2'" unfolding simps by auto from tighten_poly_bounds[OF tb1 ur(1) pc(1) sr(1)] have rc1: "root_cond (p1, l1', r1') (the_unique_root (p1, l1, r1))" and bnd1: "l1 \ l1'" "l1' \ r1'" "r1' \ r1" and d1: "r1' - l1' = (r1 - l1) / 2" and sr1: "sr1' = sgn (ipoly p1 r1')" by auto from pc have "p1 \ 0" "p2 \ 0" by auto from unique_root_sub_interval[OF ur(1) rc1 bnd1(1,3)] xy ur this have ur1: "unique_root (p1, l1', r1')" and x: "x = the_unique_root (p1, l1', r1')" by (auto intro!: the_unique_root_eqI) from tighten_poly_bounds[OF tb2 ur(2) pc(2) sr(2)] have rc2: "root_cond (p2, l2', r2') (the_unique_root (p2, l2, r2))" and bnd2: "l2 \ l2'" "l2' \ r2'" "r2' \ r2" and d2: "r2' - l2' = (r2 - l2) / 2" and sr2: "sr2' = sgn (ipoly p2 r2')" by auto from unique_root_sub_interval[OF ur(2) rc2 bnd2(1,3)] xy ur pc have ur2: "unique_root (p2, l2', r2')" and y: "y = the_unique_root (p2, l2', r2')" by auto define d' where "d' = d/2" have d': "d' = r1' - l1' + (r2' - l2')" unfolding d'_def d d1 d2 by (simp add: field_simps) have d'0: "d' \ 0" using bnd1 bnd2 unfolding d' by auto have dd: "d - d' = d/2" unfolding d'_def by simp have "abs (x - y) \ 2 * ?r d" proof (rule ccontr) assume "\ ?thesis" hence lt: "2 * ?r d < abs (x - y)" by auto have "r1 - l1 \ d" "r2 - l2 \ d" unfolding d using bnd1 bnd2 by auto from this[folded of_rat_less_eq[where 'a = real]] lt have "?r (r1 - l1) < abs (x - y) / 2" "?r (r2 - l2) < abs (x - y) / 2" and dd: "?r r1 - ?r l1 \ ?r d" "?r r2 - ?r l2 \ ?r d" by (auto simp: of_rat_diff) from le have "r1 \ l2" by auto hence r1l2: "?r r1 \ ?r l2" unfolding of_rat_less_eq by auto from False have "r2 \ l1" by auto hence r2l1: "?r r2 \ ?r l1" unfolding of_rat_less_eq by auto show False proof (cases "x \ y") case True from urx(1-2) dd(1) have "?r r1 \ x + ?r d" by auto with r1l2 have "?r l2 \ x + ?r d" by auto with True lt ury(2) dd(2) show False by auto next case False from ury(1-2) dd(2) have "?r r2 \ y + ?r d" by auto with r2l1 have "?r l1 \ y + ?r d" by auto with False lt urx(2) dd(1) show False by auto qed qed hence dd': "delta_gt delta (?r d) (?r d')" unfolding delta_gt_def delta_def using dd by (auto simp: hom_distribs) show ?thesis unfolding l by (rule IH[OF _ d' ur1 ur2 x y sr1 sr2], insert d'0 dd', auto) qed qed qed } thus ?thesis by auto qed end (* **************************************************************** *) fun real_alg_1 :: "real_alg_2 \ real_alg_1" where "real_alg_1 (Rational r) = of_rat_1 r" | "real_alg_1 (Irrational n rai) = rai" lemma real_alg_1: "real_of_1 (real_alg_1 x) = real_of_2 x" by (cases x, auto simp: of_rat_1) definition root_2 :: "nat \ real_alg_2 \ real_alg_2" where "root_2 n x = root_1 n (real_alg_1 x)" lemma root_2: assumes "invariant_2 x" shows "real_of_2 (root_2 n x) = root n (real_of_2 x)" "invariant_2 (root_2 n x)" proof (atomize(full), cases x, goal_cases) case (1 y) from of_rat_1[of y] root_1[of "of_rat_1 y" n] assms 1 real_alg_2 show ?case by (simp add: root_2_def) next case (2 i rai) from root_1[of rai n] assms 2 real_alg_2 show ?case by (auto simp: root_2_def) qed fun add_2 :: "real_alg_2 \ real_alg_2 \ real_alg_2" where "add_2 (Rational r) (Rational q) = Rational (r + q)" | "add_2 (Rational r) (Irrational n x) = Irrational n (add_rat_1 r x)" | "add_2 (Irrational n x) (Rational q) = Irrational n (add_rat_1 q x)" | "add_2 (Irrational n x) (Irrational m y) = add_1 x y" lemma add_2: assumes x: "invariant_2 x" and y: "invariant_2 y" shows "invariant_2 (add_2 x y)" (is ?g1) and "real_of_2 (add_2 x y) = real_of_2 x + real_of_2 y" (is ?g2) using assms add_rat_1 add_1 by (atomize (full), (cases x; cases y), auto simp: hom_distribs) fun mult_2 :: "real_alg_2 \ real_alg_2 \ real_alg_2" where "mult_2 (Rational r) (Rational q) = Rational (r * q)" | "mult_2 (Rational r) (Irrational n y) = mult_rat_1 r y" | "mult_2 (Irrational n x) (Rational q) = mult_rat_1 q x" | "mult_2 (Irrational n x) (Irrational m y) = mult_1 x y" lemma mult_2: assumes "invariant_2 x" "invariant_2 y" shows "real_of_2 (mult_2 x y) = real_of_2 x * real_of_2 y" "invariant_2 (mult_2 x y)" using assms by (atomize(full), (cases x; cases y; auto simp: mult_rat_1 mult_1 hom_distribs)) fun to_rat_2 :: "real_alg_2 \ rat option" where "to_rat_2 (Rational r) = Some r" | "to_rat_2 (Irrational n rai) = None" lemma to_rat_2: assumes rc: "invariant_2 x" shows "to_rat_2 x = (if real_of_2 x \ \ then Some (THE q. real_of_2 x = of_rat q) else None)" proof (cases x) case (Irrational n rai) from real_of_2_Irrational[OF rc[unfolded this]] show ?thesis unfolding Irrational Rats_def by auto qed simp fun equal_2 :: "real_alg_2 \ real_alg_2 \ bool" where "equal_2 (Rational r) (Rational q) = (r = q)" | "equal_2 (Irrational n (p,_)) (Irrational m (q,_)) = (p = q \ n = m)" | "equal_2 (Rational r) (Irrational _ yy) = False" | "equal_2 (Irrational _ xx) (Rational q) = False" lemma equal_2[simp]: assumes rc: "invariant_2 x" "invariant_2 y" shows "equal_2 x y = (real_of_2 x = real_of_2 y)" using info_2[OF rc] by (cases x; cases y, auto) fun compare_2 :: "real_alg_2 \ real_alg_2 \ order" where "compare_2 (Rational r) (Rational q) = (compare r q)" | "compare_2 (Irrational n (p,l,r)) (Irrational m (q,l',r')) = (if p = q \ n = m then Eq else compare_1 p q l r (sgn (ipoly p r)) l' r' (sgn (ipoly q r')))" | "compare_2 (Rational r) (Irrational _ xx) = (compare_rat_1 r xx)" | "compare_2 (Irrational _ xx) (Rational r) = (invert_order (compare_rat_1 r xx))" lemma compare_2: assumes rc: "invariant_2 x" "invariant_2 y" shows "compare_2 x y = compare (real_of_2 x) (real_of_2 y)" proof (cases x) case (Rational r) note xx = this show ?thesis proof (cases y) case (Rational q) note yy = this show ?thesis unfolding xx yy by (simp add: compare_rat_def compare_real_def comparator_of_def of_rat_less) next case (Irrational n yy) note yy = this from compare_rat_1 rc show ?thesis unfolding xx yy by (simp add: of_rat_1) qed next case (Irrational n xx) note xx = this show ?thesis proof (cases y) case (Rational q) note yy = this from compare_rat_1 rc show ?thesis unfolding xx yy by simp next case (Irrational m yy) note yy = this obtain p l r where xxx: "xx = (p,l,r)" by (cases xx) obtain q l' r' where yyy: "yy = (q,l',r')" by (cases yy) note rc = rc[unfolded xx xxx yy yyy] from rc have I: "invariant_1_2 (p,l,r)" "invariant_1_2 (q,l',r')" by auto then have "unique_root (p,l,r)" "unique_root (q,l',r')" "poly_cond2 p" "poly_cond2 q" by auto from compare_1[OF this _ refl refl] show ?thesis using equal_2[OF rc] unfolding xx xxx yy yyy by simp qed qed fun sgn_2 :: "real_alg_2 \ rat" where "sgn_2 (Rational r) = sgn r" | "sgn_2 (Irrational n rai) = sgn_1 rai" lemma sgn_2: "invariant_2 x \ real_of_rat (sgn_2 x) = sgn (real_of_2 x)" using sgn_1 by (cases x, auto simp: real_of_rat_sgn) fun floor_2 :: "real_alg_2 \ int" where "floor_2 (Rational r) = floor r" | "floor_2 (Irrational n rai) = floor_1 rai" lemma floor_2: "invariant_2 x \ floor_2 x = floor (real_of_2 x)" by (cases x, auto simp: floor_1) (* *************** *) subsubsection \Definitions and Algorithms on Type with Invariant\ lift_definition of_rat_3 :: "rat \ real_alg_3" is of_rat_2 by (auto simp: of_rat_2) lemma of_rat_3: "real_of_3 (of_rat_3 x) = of_rat x" by (transfer, auto simp: of_rat_2) lift_definition root_3 :: "nat \ real_alg_3 \ real_alg_3" is root_2 by (auto simp: root_2) lemma root_3: "real_of_3 (root_3 n x) = root n (real_of_3 x)" by (transfer, auto simp: root_2) lift_definition equal_3 :: "real_alg_3 \ real_alg_3 \ bool" is equal_2 . lemma equal_3: "equal_3 x y = (real_of_3 x = real_of_3 y)" by (transfer, auto) lift_definition compare_3 :: "real_alg_3 \ real_alg_3 \ order" is compare_2 . lemma compare_3: "compare_3 x y = (compare (real_of_3 x) (real_of_3 y))" by (transfer, auto simp: compare_2) lift_definition add_3 :: "real_alg_3 \ real_alg_3 \ real_alg_3" is add_2 by (auto simp: add_2) lemma add_3: "real_of_3 (add_3 x y) = real_of_3 x + real_of_3 y" by (transfer, auto simp: add_2) lift_definition mult_3 :: "real_alg_3 \ real_alg_3 \ real_alg_3" is mult_2 by (auto simp: mult_2) lemma mult_3: "real_of_3 (mult_3 x y) = real_of_3 x * real_of_3 y" by (transfer, auto simp: mult_2) lift_definition sgn_3 :: "real_alg_3 \ rat" is sgn_2 . lemma sgn_3: "real_of_rat (sgn_3 x) = sgn (real_of_3 x)" by (transfer, auto simp: sgn_2) lift_definition to_rat_3 :: "real_alg_3 \ rat option" is to_rat_2 . lemma to_rat_3: "to_rat_3 x = (if real_of_3 x \ \ then Some (THE q. real_of_3 x = of_rat q) else None)" by (transfer, simp add: to_rat_2) lift_definition floor_3 :: "real_alg_3 \ int" is floor_2 . lemma floor_3: "floor_3 x = floor (real_of_3 x)" by (transfer, auto simp: floor_2) (* *************** *) (* info *) lift_definition info_3 :: "real_alg_3 \ rat + int poly \ nat" is info_2 . lemma info_3_fun: "real_of_3 x = real_of_3 y \ info_3 x = info_3 y" by (transfer, intro info_2_unique, auto) lift_definition info_real_alg :: "real_alg \ rat + int poly \ nat" is info_3 by (metis info_3_fun) lemma info_real_alg: "info_real_alg x = Inr (p,n) \ p represents (real_of x) \ card {y. y \ real_of x \ ipoly p y = 0} = n \ irreducible p" "info_real_alg x = Inl q \ real_of x = of_rat q" proof (atomize(full), transfer, transfer, goal_cases) case (1 x p n q) from 1 have x: "invariant_2 x" by auto note info = info_2_card[OF this] show ?case proof (cases x) case irr: (Irrational m rai) from info(1)[of p n] show ?thesis unfolding irr by (cases rai, auto simp: poly_cond_def) qed (insert 1 info, auto) qed (* add *) instantiation real_alg :: plus begin lift_definition plus_real_alg :: "real_alg \ real_alg \ real_alg" is add_3 by (simp add: add_3) instance .. end lemma plus_real_alg: "(real_of x) + (real_of y) = real_of (x + y)" by (transfer, rule add_3[symmetric]) (* minus *) instantiation real_alg :: minus begin definition minus_real_alg :: "real_alg \ real_alg \ real_alg" where "minus_real_alg x y = x + (-y)" instance .. end lemma minus_real_alg: "(real_of x) - (real_of y) = real_of (x - y)" unfolding minus_real_alg_def minus_real_def uminus_real_alg plus_real_alg .. (* of_rat *) lift_definition of_rat_real_alg :: "rat \ real_alg" is of_rat_3 . lemma of_rat_real_alg: "real_of_rat x = real_of (of_rat_real_alg x)" by (transfer, rule of_rat_3[symmetric]) (* zero *) instantiation real_alg :: zero begin definition zero_real_alg :: "real_alg" where "zero_real_alg \ of_rat_real_alg 0" instance .. end lemma zero_real_alg: "0 = real_of 0" unfolding zero_real_alg_def by (simp add: of_rat_real_alg[symmetric]) (* one *) instantiation real_alg :: one begin definition one_real_alg :: "real_alg" where "one_real_alg \ of_rat_real_alg 1" instance .. end lemma one_real_alg: "1 = real_of 1" unfolding one_real_alg_def by (simp add: of_rat_real_alg[symmetric]) (* times *) instantiation real_alg :: times begin lift_definition times_real_alg :: "real_alg \ real_alg \ real_alg" is mult_3 by (simp add: mult_3) instance .. end lemma times_real_alg: "(real_of x) * (real_of y) = real_of (x * y)" by (transfer, rule mult_3[symmetric]) (* inverse *) instantiation real_alg :: inverse begin lift_definition inverse_real_alg :: "real_alg \ real_alg" is inverse_3 by (simp add: inverse_3) definition divide_real_alg :: "real_alg \ real_alg \ real_alg" where "divide_real_alg x y = x * inverse y" (* TODO: better to use poly_div *) instance .. end lemma inverse_real_alg: "inverse (real_of x) = real_of (inverse x)" by (transfer, rule inverse_3[symmetric]) lemma divide_real_alg: "(real_of x) / (real_of y) = real_of (x / y)" unfolding divide_real_alg_def times_real_alg[symmetric] divide_real_def inverse_real_alg .. (* group *) instance real_alg :: ab_group_add apply intro_classes apply (transfer, unfold add_3, force) apply (unfold zero_real_alg_def, transfer, unfold add_3 of_rat_3, force) apply (transfer, unfold add_3 of_rat_3, force) apply (transfer, unfold add_3 uminus_3 of_rat_3, force) apply (unfold minus_real_alg_def, force) done (* field *) instance real_alg :: field apply intro_classes apply (transfer, unfold mult_3, force) apply (transfer, unfold mult_3, force) apply (unfold one_real_alg_def, transfer, unfold mult_3 of_rat_3, force) apply (transfer, unfold mult_3 add_3, force simp: field_simps) apply (unfold zero_real_alg_def, transfer, unfold of_rat_3, force) apply (transfer, unfold mult_3 inverse_3 of_rat_3, force simp: field_simps) apply (unfold divide_real_alg_def, force) apply (transfer, unfold inverse_3 of_rat_3, force) done (* numeral *) instance real_alg :: numeral .. (* root *) lift_definition root_real_alg :: "nat \ real_alg \ real_alg" is root_3 by (simp add: root_3) lemma root_real_alg: "root n (real_of x) = real_of (root_real_alg n x)" by (transfer, rule root_3[symmetric]) (* sgn *) lift_definition sgn_real_alg_rat :: "real_alg \ rat" is sgn_3 by (insert sgn_3, metis to_rat_of_rat) lemma sgn_real_alg_rat: "real_of_rat (sgn_real_alg_rat x) = sgn (real_of x)" by (transfer, auto simp: sgn_3) instantiation real_alg :: sgn begin definition sgn_real_alg :: "real_alg \ real_alg" where "sgn_real_alg x = of_rat_real_alg (sgn_real_alg_rat x)" instance .. end lemma sgn_real_alg: "sgn (real_of x) = real_of (sgn x)" unfolding sgn_real_alg_def of_rat_real_alg[symmetric] by (transfer, simp add: sgn_3) (* equal *) instantiation real_alg :: equal begin lift_definition equal_real_alg :: "real_alg \ real_alg \ bool" is equal_3 by (simp add: equal_3) instance proof fix x y :: real_alg show "equal_class.equal x y = (x = y)" by (transfer, simp add: equal_3) qed end lemma equal_real_alg: "HOL.equal (real_of x) (real_of y) = (x = y)" unfolding equal_real_def by (transfer, auto) (* comparisons *) instantiation real_alg :: ord begin definition less_real_alg :: "real_alg \ real_alg \ bool" where [code del]: "less_real_alg x y = (real_of x < real_of y)" definition less_eq_real_alg :: "real_alg \ real_alg \ bool" where [code del]: "less_eq_real_alg x y = (real_of x \ real_of y)" instance .. end lemma less_real_alg: "less (real_of x) (real_of y) = (x < y)" unfolding less_real_alg_def .. lemma less_eq_real_alg: "less_eq (real_of x) (real_of y) = (x \ y)" unfolding less_eq_real_alg_def .. instantiation real_alg :: compare_order begin lift_definition compare_real_alg :: "real_alg \ real_alg \ order" is compare_3 by (simp add: compare_3) lemma compare_real_alg: "compare (real_of x) (real_of y) = (compare x y)" by (transfer, simp add: compare_3) instance proof (intro_classes, unfold compare_real_alg[symmetric, abs_def]) show "le_of_comp (\x y. compare (real_of x) (real_of y)) = (\)" by (intro ext, auto simp: compare_real_def comparator_of_def le_of_comp_def less_eq_real_alg_def) show "lt_of_comp (\x y. compare (real_of x) (real_of y)) = (<)" by (intro ext, auto simp: compare_real_def comparator_of_def lt_of_comp_def less_real_alg_def) show "comparator (\x y. compare (real_of x) (real_of y))" unfolding comparator_def proof (intro conjI impI allI) fix x y z :: "real_alg" let ?r = real_of note rc = comparator_compare[where 'a = real, unfolded comparator_def] from rc show "invert_order (compare (?r x) (?r y)) = compare (?r y) (?r x)" by blast from rc show "compare (?r x) (?r y) = Lt \ compare (?r y) (?r z) = Lt \ compare (?r x) (?r z) = Lt" by blast assume "compare (?r x) (?r y) = Eq" with rc have "?r x = ?r y" by blast thus "x = y" unfolding real_of_inj . qed qed end lemma less_eq_real_alg_code[code]: "(less_eq :: real_alg \ real_alg \ bool) = le_of_comp compare" "(less :: real_alg \ real_alg \ bool) = lt_of_comp compare" by (rule ord_defs(1)[symmetric], rule ord_defs(2)[symmetric]) instantiation real_alg :: abs begin definition abs_real_alg :: "real_alg \ real_alg" where "abs_real_alg x = (if real_of x < 0 then uminus x else x)" instance .. end lemma abs_real_alg: "abs (real_of x) = real_of (abs x)" unfolding abs_real_alg_def abs_real_def if_distrib by (auto simp: uminus_real_alg) lemma sgn_real_alg_sound: "sgn x = (if x = 0 then 0 else if 0 < real_of x then 1 else - 1)" (is "_ = ?r") proof - have "real_of (sgn x) = sgn (real_of x)" by (simp add: sgn_real_alg) also have "\ = real_of ?r" unfolding sgn_real_def if_distrib by (auto simp: less_real_alg_def zero_real_alg_def one_real_alg_def of_rat_real_alg[symmetric] equal_real_alg[symmetric] equal_real_def uminus_real_alg[symmetric]) finally show "sgn x = ?r" unfolding equal_real_alg[symmetric] equal_real_def by simp qed lemma real_of_of_int: "real_of_rat (rat_of_int z) = real_of (of_int z)" proof (cases "z \ 0") case True define n where "n = nat z" from True have z: "z = int n" unfolding n_def by simp show ?thesis unfolding z by (induct n, auto simp: zero_real_alg plus_real_alg[symmetric] one_real_alg hom_distribs) next case False define n where "n = nat (-z)" from False have z: "z = - int n" unfolding n_def by simp show ?thesis unfolding z by (induct n, auto simp: zero_real_alg plus_real_alg[symmetric] one_real_alg uminus_real_alg[symmetric] minus_real_alg[symmetric] hom_distribs) qed instance real_alg :: linordered_field apply standard apply (unfold less_eq_real_alg_def plus_real_alg[symmetric], force) apply (unfold abs_real_alg_def less_real_alg_def zero_real_alg[symmetric], rule refl) apply (unfold less_real_alg_def times_real_alg[symmetric], force) apply (rule sgn_real_alg_sound) done instantiation real_alg :: floor_ceiling begin lift_definition floor_real_alg :: "real_alg \ int" is floor_3 by (auto simp: floor_3) lemma floor_real_alg: "floor (real_of x) = floor x" by (transfer, auto simp: floor_3) instance proof fix x :: real_alg show "of_int \x\ \ x \ x < of_int (\x\ + 1)" unfolding floor_real_alg[symmetric] using floor_correct[of "real_of x"] unfolding less_eq_real_alg_def less_real_alg_def real_of_of_int[symmetric] by (auto simp: hom_distribs) hence "x \ of_int (\x\ + 1)" by auto thus "\z. x \ of_int z" by blast qed end instantiation real_alg :: "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}" begin definition [simp]: "normalize_real_alg = (normalize_field :: real_alg \ _)" definition [simp]: "unit_factor_real_alg = (unit_factor_field :: real_alg \ _)" definition [simp]: "modulo_real_alg = (mod_field :: real_alg \ _)" definition [simp]: "euclidean_size_real_alg = (euclidean_size_field :: real_alg \ _)" definition [simp]: "division_segment (x :: real_alg) = 1" instance by standard (simp_all add: dvd_field_iff field_split_simps split: if_splits) end instantiation real_alg :: euclidean_ring_gcd begin definition gcd_real_alg :: "real_alg \ real_alg \ real_alg" where "gcd_real_alg = Euclidean_Algorithm.gcd" definition lcm_real_alg :: "real_alg \ real_alg \ real_alg" where "lcm_real_alg = Euclidean_Algorithm.lcm" definition Gcd_real_alg :: "real_alg set \ real_alg" where "Gcd_real_alg = Euclidean_Algorithm.Gcd" definition Lcm_real_alg :: "real_alg set \ real_alg" where "Lcm_real_alg = Euclidean_Algorithm.Lcm" instance by standard (simp_all add: gcd_real_alg_def lcm_real_alg_def Gcd_real_alg_def Lcm_real_alg_def) end instance real_alg :: field_gcd .. definition min_int_poly_real_alg :: "real_alg \ int poly" where "min_int_poly_real_alg x = (case info_real_alg x of Inl r \ poly_rat r | Inr (p,_) \ p)" lemma min_int_poly_real_alg_real_of: "min_int_poly_real_alg x = min_int_poly (real_of x)" proof (cases "info_real_alg x") case (Inl r) show ?thesis unfolding info_real_alg(2)[OF Inl] min_int_poly_real_alg_def Inl by (simp add: min_int_poly_of_rat) next case (Inr pair) then obtain p n where Inr: "info_real_alg x = Inr (p,n)" by (cases pair, auto) hence "poly_cond p" by (transfer, transfer, auto simp: info_2_card) hence "min_int_poly (real_of x) = p" using info_real_alg(1)[OF Inr] by (intro min_int_poly_unique, auto) thus ?thesis unfolding min_int_poly_real_alg_def Inr by simp qed lemma min_int_poly_real_code: "min_int_poly_real (real_of x) = min_int_poly_real_alg x" by (simp add: min_int_poly_real_alg_real_of) lemma min_int_poly_real_of: "min_int_poly (real_of x) = min_int_poly x" proof (rule min_int_poly_unique[OF _ min_int_poly_irreducible lead_coeff_min_int_poly_pos]) show "min_int_poly x represents real_of x" oops (* TODO: this gives an implementation of min-int-poly on type real-alg qed lemma min_int_poly_real_alg_code_unfold [code_unfold]: "min_int_poly = min_int_poly_real_alg" by (intro ext, unfold min_int_poly_real_alg_real_of, simp add: min_int_poly_real_of) *) definition real_alg_of_real :: "real \ real_alg" where "real_alg_of_real x = (if (\ y. x = real_of y) then (THE y. x = real_of y) else 0)" lemma real_alg_of_real_code[code]: "real_alg_of_real (real_of x) = x" using real_of_inj unfolding real_alg_of_real_def by auto lift_definition to_rat_real_alg_main :: "real_alg \ rat option" is to_rat_3 by (simp add: to_rat_3) lemma to_rat_real_alg_main: "to_rat_real_alg_main x = (if real_of x \ \ then Some (THE q. real_of x = of_rat q) else None)" by (transfer, simp add: to_rat_3) definition to_rat_real_alg :: "real_alg \ rat" where "to_rat_real_alg x = (case to_rat_real_alg_main x of Some q \ q | None \ 0)" definition is_rat_real_alg :: "real_alg \ bool" where "is_rat_real_alg x = (case to_rat_real_alg_main x of Some q \ True | None \ False)" lemma is_rat_real_alg: "is_rat (real_of x) = (is_rat_real_alg x)" unfolding is_rat_real_alg_def is_rat to_rat_real_alg_main by auto lemma to_rat_real_alg: "to_rat (real_of x) = (to_rat_real_alg x)" unfolding to_rat to_rat_real_alg_def to_rat_real_alg_main by auto +definition algebraic_real :: "real \ bool" where + [simp]: "algebraic_real = algebraic" + +lemma algebraic_real_iff[code_unfold]: "algebraic = algebraic_real" by simp + +lemma algebraic_real_code[code]: "algebraic_real (real_of x) = True" +proof (cases "info_real_alg x") + case (Inl r) + show ?thesis using info_real_alg(2)[OF Inl] by (auto simp: algebraic_of_rat) +next + case (Inr pair) + then obtain p n where Inr: "info_real_alg x = Inr (p,n)" by (cases pair, auto) + from info_real_alg(1)[OF Inr] have "p represents (real_of x)" by auto + thus ?thesis by (auto simp: algebraic_altdef_ipoly) +qed + subsection \Real Algebraic Numbers as Implementation for Real Numbers\ lemmas real_alg_code_eqns = one_real_alg zero_real_alg uminus_real_alg root_real_alg minus_real_alg plus_real_alg times_real_alg inverse_real_alg divide_real_alg equal_real_alg less_real_alg less_eq_real_alg compare_real_alg sgn_real_alg abs_real_alg floor_real_alg is_rat_real_alg to_rat_real_alg min_int_poly_real_code code_datatype real_of declare [[code drop: "plus :: real \ real \ real" "uminus :: real \ real" "minus :: real \ real \ real" "times :: real \ real \ real" "inverse :: real \ real" "divide :: real \ real \ real" "floor :: real \ int" "HOL.equal :: real \ real \ bool" "compare :: real \ real \ order" "less_eq :: real \ real \ bool" "less :: real \ real \ bool" "0 :: real" "1 :: real" "sgn :: real \ real" "abs :: real \ real" min_int_poly_real root]] declare real_alg_code_eqns [code equation] lemma Ratreal_code[code]: "Ratreal = real_of \ of_rat_real_alg" by (transfer, transfer) (simp add: fun_eq_iff of_rat_2) lemma real_of_post[code_post]: "real_of (Real_Alg_Quotient (Real_Alg_Invariant (Rational x))) = of_rat x" proof (transfer) fix x show "real_of_3 (Real_Alg_Invariant (Rational x)) = real_of_rat x" by (simp add: Real_Alg_Invariant_inverse real_of_3.rep_eq) qed end diff --git a/thys/Algebraic_Numbers/Real_Factorization.thy b/thys/Algebraic_Numbers/Real_Factorization.thy deleted file mode 100644 --- a/thys/Algebraic_Numbers/Real_Factorization.thy +++ /dev/null @@ -1,282 +0,0 @@ -section \Real Factorization\ - -text \This theory contains an algorithm to completely factorize real polynomials with rational - coefficients. It internally does a complex polynomial factorization, and then combines - all non-real roots with their conjugates.\ - -theory Real_Factorization -imports - Complex_Algebraic_Numbers -begin - -fun delete_cnj :: "complex \ nat \ (complex \ nat) list \ (complex \ nat) list" where - "delete_cnj x i ((y,j) # yjs) = (if x = y then if Suc j = i then yjs else if Suc j > i then - ((y,j - i) # yjs) else delete_cnj x (i - Suc j) yjs else (y,j) # delete_cnj x i yjs)" -| "delete_cnj _ _ [] = []" - -lemma delete_cnj_length[termination_simp]: "length (delete_cnj x i yjs) \ length yjs" - by (induct x i yjs rule: delete_cnj.induct, auto) - -fun complex_roots_to_real_factorization :: "(complex \ nat) list \ (real poly \ nat)list" where - "complex_roots_to_real_factorization [] = []" -| "complex_roots_to_real_factorization ((x,i) # xs) = (if x \ \ then - ([:-(Re x),1:],Suc i) # complex_roots_to_real_factorization xs else - let xx = cnj x; ys = delete_cnj xx (Suc i) xs; p = map_poly Re ([:-x,1:] * [:-xx,1:]) - in (p,Suc i) # complex_roots_to_real_factorization ys)" - -definition factorize_real_poly :: "real poly \ (real \ (real poly \ nat) list) option" where - "factorize_real_poly p \ map_option - (\ (c,ris). (Re c, complex_roots_to_real_factorization ris)) - (factorize_complex_main (map_poly of_real p))" - - -lemma monic_imp_nonzero: "monic x \ x \ 0" for x :: "'a :: semiring_1 poly" by auto - -lemma delete_cnj: assumes - "order x (\(x, i)\xis. [:- x, 1:] ^ Suc i) \ si" "si \ 0" - shows "(\(x, i)\xis. [:- x, 1:] ^ Suc i) = - [:- x, 1:] ^ si * (\(x, i)\delete_cnj x si xis. [:- x, 1:] ^ Suc i)" -using assms -proof (induct x si xis rule: delete_cnj.induct) - case (2 x si) - hence "order x 1 \ 1" by auto - hence "[:-x,1:]^1 dvd 1" unfolding order_divides by simp - from power_le_dvd[OF this, of 1] \si \ 0\ have "[:- x, 1:] dvd 1" by simp - from divides_degree[OF this] - show ?case by auto -next - case (1 x i y j yjs) - note IH = 1(1-2) - let ?yj = "[:-y,1:]^Suc j" - let ?yjs = "(\(x,i)\yjs. [:- x, 1:] ^ Suc i)" - let ?x = "[: - x, 1 :]" - let ?xi = "?x ^ i" - have "monic (\(x,i)\(y, j) # yjs. [:- x, 1:] ^ Suc i)" - by (rule monic_prod_list_pow) then have "monic (?yj * ?yjs)" by simp - from monic_imp_nonzero[OF this] have yy0: "?yj * ?yjs \ 0" by auto - have id: "(\(x,i)\(y, j) # yjs. [:- x, 1:] ^ Suc i) = ?yj * ?yjs" by simp - from 1(3-) have ord: "i \ order x (?yj * ?yjs)" and i: "i \ 0" unfolding id by auto - from ord[unfolded order_mult[OF yy0]] have ord: "i \ order x ?yj + order x ?yjs" . - from this[unfolded order_linear_power'] - have ord: "i \ (if y = x then Suc j else 0) + order x ?yjs" by simp - show ?case - proof (cases "x = y") - case False - from ord False have "i \ order x ?yjs" by simp - note IH = IH(2)[OF False this i] - from False have del: "delete_cnj x i ((y, j) # yjs) = (y,j) # delete_cnj x i yjs" by simp - show ?thesis unfolding del id IH - by (simp add: ac_simps) - next - case True note xy = this - note IH = IH(1)[OF True] - show ?thesis - proof (cases "Suc j \ i") - case False - from ord have ord: "i - Suc j \ order x ?yjs" unfolding xy by simp - have "?xi = ?x ^ (Suc j + (i - Suc j))" using False by simp - also have "\ = ?x ^ Suc j * ?x ^ (i - Suc j)" - unfolding power_add by simp - finally have xi: "?xi = ?x ^ Suc j * ?x ^ (i - Suc j)" . - from False have "Suc j \ i" "\ i < Suc j" "i - Suc j \ 0" by auto - note IH = IH[OF this(1,2) ord this(3)] - from xy False have del: "delete_cnj x i ((y, j) # yjs) = delete_cnj x (i - Suc j) yjs" by auto - show ?thesis unfolding del id unfolding IH xi unfolding xy by simp - next - case True - hence "Suc j = i \ i < Suc j" by auto - thus ?thesis - proof - assume i: "Suc j = i" - from xy i have del: "delete_cnj x i ((y, j) # yjs) = yjs" by simp - show ?thesis unfolding id del unfolding xy i by simp - next - assume ij: "i < Suc j" - with xy i have del: "delete_cnj x i ((y, j) # yjs) = (y, j - i) # yjs" by simp - from ij have idd: "Suc j = i + Suc (j - i)" by simp - show ?thesis unfolding id del unfolding xy idd power_add by simp - qed - qed - qed -qed - - -lemma factorize_real_poly: assumes fp: "factorize_real_poly p = Some (c,qis)" - shows - "p = smult c (\(q, i)\qis. q ^ i)" - "(q,j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1,2}" -proof - - interpret map_poly_inj_idom_hom of_real.. - have "(p = smult c (\(q, i)\qis. q ^ i)) \ ((q,j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1,2})" - proof (cases "p = 0") - case True - have "factorize_real_poly p = Some (0,[])" unfolding True - by (simp add: factorize_real_poly_def factorize_complex_main_def yun_factorization_def) - with fp have id: "c = 0" "qis = []" by auto - thus ?thesis unfolding True by simp - next - case False note p0 = this - let ?c = complex_of_real - let ?rp = "map_poly Re" - let ?cp = "map_poly ?c" - let ?p = "?cp p" - from fp[unfolded factorize_real_poly_def] - obtain d xis where fp: "factorize_complex_main ?p = Some (d,xis)" - and c: "c = Re d" and qis: "qis = complex_roots_to_real_factorization xis" by auto - from factorize_complex_main[OF fp] have p: "?p = smult d (\(x, i)\xis. [:- x, 1:] ^ Suc i)" - (is "_ = smult d ?q") . - from arg_cong[OF this, of "\ p. coeff p (degree p)"] - have "coeff ?p (degree ?p) = coeff (smult d ?q) (degree (smult d ?q))" . - also have "coeff ?p (degree ?p) = ?c (coeff p (degree p))" by simp - also have "coeff (smult d ?q) (degree (smult d ?q)) = d * coeff ?q (degree ?q)" - by simp - also have "monic ?q" by (rule monic_prod_list_pow) - finally have d: "d = ?c (coeff p (degree p))" by auto - from arg_cong[OF this, of Re, folded c] have c: "c = coeff p (degree p)" by auto - have "set (coeffs ?p) \ \" by auto - with p have q': "set (coeffs (smult d ?q)) \ \" by auto - from d p0 have d0: "d \ 0" by auto - have "smult d ?q = [:d:] * ?q" by auto - from real_poly_factor[OF q'[unfolded this]] d0 d - have q: "set (coeffs ?q) \ \" by auto - have "p = ?rp ?p" - by (rule sym, subst map_poly_map_poly, force, rule map_poly_idI, auto) - also have "\ = ?rp (smult d ?q)" unfolding p .. - also have "?q = ?cp (?rp ?q)" - by (rule sym, rule map_poly_of_real_Re, insert q, auto) - also have "d = ?c c" unfolding d c .. - also have "smult (?c c) (?cp (?rp ?q)) = ?cp (smult c (?rp ?q))" by (simp add: hom_distribs) - also have "?rp \ = smult c (?rp ?q)" - by (subst map_poly_map_poly, force, rule map_poly_idI, auto) - finally have p: "p = smult c (?rp ?q)" . - let ?fact = complex_roots_to_real_factorization - have "?rp ?q = (\(q, i)\qis. q ^ i) \ - ((q, j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1, 2})" - using q unfolding qis - proof (induct xis rule: complex_roots_to_real_factorization.induct) - case 1 - show ?case by simp - next - case (2 x i xis) - note IH = 2(1-2) - note prems = 2(3) - let ?xi = "[:- x, 1:] ^ Suc i" - let ?xis = "(\(x, i)\xis. [:- x, 1:] ^ Suc i)" - have id: "(\(x, i)\((x,i) # xis). [:- x, 1:] ^ Suc i) = ?xi * ?xis" - by simp - show ?case - proof (cases "x \ \") - case True - have xi: "set (coeffs ?xi) \ \" - by (rule real_poly_power, insert True, auto) - have xis: "set (coeffs ?xis) \ \" - by (rule real_poly_factor[OF prems[unfolded id] xi], rule linear_power_nonzero) - note IH = IH(1)[OF True xis] - have "?rp (?xi * ?xis) = ?rp ?xi * ?rp ?xis" - by (rule map_poly_Re_mult[OF xi xis]) - also have "?rp ?xi = (?rp [: -x,1 :])^Suc i" - by (rule map_poly_Re_power, insert True, auto) - also have "?rp [: -x,1 :] = [:-(Re x),1:]" by auto - also have "?rp ?xis = (\ (a,b) \ ?fact xis. a ^ b)" - using IH by auto - also have "[:- Re x, 1:] ^ Suc i * (\ (a,b) \ ?fact xis. a ^ b) = - (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" using True by simp - finally have idd: "?rp (?xi * ?xis) = (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" . - show ?thesis unfolding id idd - proof (intro conjI, force, intro impI) - assume "(q, j) \ set (?fact ((x, i) # xis))" - hence "(q,j) \ set (?fact xis) \ (q = [:- Re x, 1:] \ j = Suc i)" - using True by auto - thus "irreducible q \ j \ 0 \ monic q \ degree q \ {1, 2}" - proof - assume "(q,j) \ set (?fact xis)" - with IH show ?thesis by auto - next - assume "q = [:- Re x, 1:] \ j = Suc i" - with linear_irreducible_field[of "[:- Re x, 1:]"] show ?thesis by auto - qed - qed - next - case False - define xi where "xi = [:Re x * Re x + Im x * Im x, - (2 * Re x), 1:]" - obtain xx where xx: "xx = cnj x" by auto - have xi: "xi = ?rp ([:-x,1:] * [:-xx,1:])" unfolding xx xi_def by auto - have cpxi: "?cp xi = [:-x,1:] * [:-xx,1:]" unfolding xi_def - by (cases x, auto simp: xx legacy_Complex_simps) - obtain yis where yis: "yis = delete_cnj xx (Suc i) xis" by auto - from False have fact: "?fact ((x,i) # xis) = ((xi,Suc i) # ?fact yis)" - unfolding xi_def xx yis by simp - note IH = IH(2)[OF False xx yis xi] - have "irreducible xi" - apply (fold irreducible_connect_field) - proof (rule irreducible\<^sub>dI) - show "degree xi > 0" unfolding xi by auto - fix q p :: "real poly" - assume "degree q > 0" "degree q < degree xi" and qp: "xi = q * p" - hence dq: "degree q = 1" unfolding xi by auto - have dxi: "degree xi = 2" "xi \ 0" unfolding xi by auto - with qp have "q \ 0" "p \ 0" by auto - hence "degree xi = degree q + degree p" unfolding qp - by (rule degree_mult_eq) - with dq have dp: "degree p = 1" unfolding dxi by simp - { - fix c :: complex - assume rt: "poly (?cp xi) c = 0" - hence "poly (?cp q * ?cp p) c = 0" by (simp add: qp hom_distribs) - hence "(poly (?cp q) c = 0 \ poly (?cp p) c = 0)" by auto - hence "c = roots1 (?cp q) \ c = roots1 (?cp p)" - using roots1[of "?cp q"] roots1[of "?cp p"] dp dq by auto - hence "c \ \" unfolding roots1_def by auto - hence "c \ x" using False by auto - } - hence "poly (?cp xi) x \ 0" by auto - thus False unfolding cpxi by simp - qed - hence xi': "irreducible xi" "monic xi" "degree xi = 2" - unfolding xi by auto - let ?xxi = "[:- xx, 1:] ^ Suc i" - let ?yis = "(\(x, i)\yis. [:- x, 1:] ^ Suc i)" - let ?yi = "(?cp xi)^Suc i" - have yi: "set (coeffs ?yi) \ \" - by (rule real_poly_power, auto simp: xi) - have mon: "monic (\(x, i)\(x, i) # xis. [:- x, 1:] ^ Suc i)" - by (rule monic_prod_list_pow) - from monic_imp_nonzero[OF this] have xixis: "?xi * ?xis \ 0" unfolding id by auto - from False have xxx: "xx \ x" unfolding xx by (cases x, auto simp: legacy_Complex_simps Reals_def) - from prems[unfolded id] have prems: "set (coeffs (?xi * ?xis)) \ \" . - from id have "[:- x, 1:] ^ Suc i dvd ?xi * ?xis" by auto - from xixis this[unfolded order_divides] - have "order x (?xi * ?xis) \ Suc i" by auto - with complex_conjugate_order[OF prems xixis, of x, folded xx] - have "order xx (?xi * ?xis) \ Suc i" by auto - hence "order xx ?xi + order xx ?xis \ Suc i" unfolding order_mult[OF xixis] . - also have "order xx ?xi = 0" unfolding order_linear_power' using xxx by simp - finally have "order xx ?xis \ Suc i" by simp - hence yis: "?xis = ?xxi * ?yis" unfolding yis - by (rule delete_cnj, simp) - hence "?xi * ?xis = (?xi * ?xxi) * ?yis" by (simp only: ac_simps) - also have "?xi * ?xxi = ([:- x, 1:] * [:- xx, 1:])^Suc i" - by (metis power_mult_distrib) - also have "[:- x, 1:] * [:- xx, 1:] = ?cp xi" unfolding cpxi .. - finally have idd: "?xi * ?xis = (?cp xi)^Suc i * ?yis" by simp - from prems[unfolded idd] have R: "set (coeffs ((?cp xi)^Suc i * ?yis)) \ \" . - have yis: "set (coeffs ?yis) \ \" - by (rule real_poly_factor[OF R yi], auto, auto simp: xi_def) - note IH = IH[OF yis] - have "?rp (?xi * ?xis) = ?rp ?yi * ?rp ?yis" unfolding idd - by (rule map_poly_Re_mult[OF yi yis]) - also have "?rp ?yi = xi^Suc i" by (fold hom_distribs, rule map_poly_Re_of_real) - also have "?rp ?yis = (\ (a,b) \ ?fact yis. a ^ b)" - using IH by auto - also have "xi ^ Suc i * (\ (a,b) \ ?fact yis. a ^ b) = - (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" unfolding fact by simp - finally have idd: "?rp (?xi * ?xis) = (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" . - show ?thesis unfolding id idd fact using IH xi' by auto - qed - qed - thus ?thesis unfolding p by simp - qed - thus "p = smult c (\(q, i)\qis. q ^ i)" - "(q,j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1,2}" by blast+ -qed -end diff --git a/thys/Algebraic_Numbers/Real_Roots.thy b/thys/Algebraic_Numbers/Real_Roots.thy --- a/thys/Algebraic_Numbers/Real_Roots.thy +++ b/thys/Algebraic_Numbers/Real_Roots.thy @@ -1,746 +1,666 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Real Roots\ text \This theory contains an algorithm to determine the set of real roots of a - rational polynomial. It further contains an algorithm which tries to determine the - real roots of real-valued polynomial, which incorporates Yun-factorization and - closed formulas for polynomials of degree 2.\ + rational polynomial. For polynomials with real coefficients, we refer to + the AFP entry "Factor-Algebraic-Polynomial".\ theory Real_Roots imports Real_Algebraic_Numbers begin hide_const (open) UnivPoly.coeff hide_const (open) Module.smult text \Division of integers, rounding to the upper value.\ definition div_ceiling :: "int \ int \ int" where "div_ceiling x y = (let q = x div y in if q * y = x then q else q + 1)" definition root_bound :: "int poly \ rat" where "root_bound p \ let n = degree p; m = 1 + div_ceiling (max_list_non_empty (map (\i. abs (coeff p i)) [0.. \round to the next higher number \2^n\, so that bisection will\ \ \stay on integers for as long as possible\ in of_int (2 ^ (log_ceiling 2 m))" partial_function (tailrec) roots_of_2_main :: "int poly \ root_info \ (rat \ rat \ nat) \ (rat \ rat)list \ real_alg_2 list \ real_alg_2 list" where [code]: "roots_of_2_main p ri cr lrs rais = (case lrs of Nil \ rais | (l,r) # lrs \ let c = cr l r in if c = 0 then roots_of_2_main p ri cr lrs rais else if c = 1 then roots_of_2_main p ri cr lrs (real_alg_2'' ri p l r # rais) else let m = (l + r) / 2 in roots_of_2_main p ri cr ((m,r) # (l,m) # lrs) rais)" definition roots_of_2_irr :: "int poly \ real_alg_2 list" where "roots_of_2_irr p = (if degree p = 1 then [Rational (Rat.Fract (- coeff p 0) (coeff p 1)) ] else let ri = root_info p; cr = root_info.l_r ri; B = root_bound p in (roots_of_2_main p ri cr [(-B,B)] []))" lemma root_imp_deg_nonzero: assumes "p \ 0" "poly p x = 0" shows "degree p \ 0" proof assume "degree p = 0" from degree0_coeffs[OF this] assms show False by auto qed lemma cauchy_root_bound: fixes x :: "'a :: real_normed_field" assumes x: "poly p x = 0" and p: "p \ 0" shows "norm x \ 1 + max_list_non_empty (map (\ i. norm (coeff p i)) [0 ..< degree p]) / norm (lead_coeff p)" (is "_ \ _ + ?max / ?nlc") proof - let ?n = "degree p" let ?p = "coeff p" let ?lc = "lead_coeff p" define ml where "ml = ?max / ?nlc" from p have lc: "?lc \ 0" by auto hence nlc: "norm ?lc > 0" by auto from root_imp_deg_nonzero[OF p x] have *: "0 \ set [0 ..< degree p]" by auto have "0 \ norm (?p 0)" by simp also have "\ \ ?max" by (rule max_list_non_empty, insert *, auto) finally have max0: "?max \ 0" . with nlc have ml0: "ml \ 0" unfolding ml_def by auto hence easy: "norm x \ 1 \ ?thesis" unfolding ml_def[symmetric] by auto show ?thesis proof (cases "norm x \ 1") case True thus ?thesis using easy by auto next case False hence nx: "norm x > 1" by simp hence x0: "x \ 0" by auto hence xn0: "0 < norm x ^ ?n" by auto from x[unfolded poly_altdef] have "x ^ ?n * ?lc = x ^ ?n * ?lc - (\i\?n. x ^ i * ?p i)" unfolding poly_altdef by (simp add: ac_simps) also have "(\i\?n. x ^ i * ?p i) = x ^ ?n * ?lc + (\i < ?n. x ^ i * ?p i)" by (subst sum.remove[of _ ?n], auto intro: sum.cong) finally have "x ^ ?n * ?lc = - (\i < ?n. x ^ i * ?p i)" by simp with lc have "x ^ ?n = - (\i < ?n. x ^ i * ?p i) / ?lc" by (simp add: field_simps) from arg_cong[OF this, of norm] have "norm x ^ ?n = norm ((\i < ?n. x ^ i * ?p i) / ?lc)" unfolding norm_power by simp also have "(\i < ?n. x ^ i * ?p i) / ?lc = (\i < ?n. x ^ i * ?p i / ?lc)" by (rule sum_divide_distrib) also have "norm \ \ (\i < ?n. norm (x ^ i * (?p i / ?lc)))" by (simp add: field_simps, rule norm_sum) also have "\ = (\i < ?n. norm x ^ i * norm (?p i / ?lc))" unfolding norm_mult norm_power .. also have "\ \ (\i < ?n. norm x ^ i * ml)" proof (rule sum_mono) fix i assume "i \ {.. norm x ^ i * ml" proof (rule mult_left_mono) show "0 \ norm x ^ i" using nx by auto show "norm (?p i / ?lc) \ ml" unfolding norm_divide ml_def by (rule divide_right_mono[OF max_list_non_empty], insert nlc i, auto) qed qed also have "\ = ml * (\i < ?n. norm x ^ i)" unfolding sum_distrib_right[symmetric] by simp also have "(\i < ?n. norm x ^ i) = (norm x ^ ?n - 1) / (norm x - 1)" by (rule geometric_sum, insert nx, auto) finally have "norm x ^ ?n \ ml * (norm x ^ ?n - 1) / (norm x - 1)" by simp from mult_left_mono[OF this, of "norm x - 1"] have "(norm x - 1) * (norm x ^ ?n) \ ml * (norm x ^ ?n - 1)" using nx by auto also have "\ = (ml * (1 - 1 / (norm x ^ ?n))) * norm x ^ ?n" using nx False x0 by (simp add: field_simps) finally have "(norm x - 1) * (norm x ^ ?n) \ (ml * (1 - 1 / (norm x ^ ?n))) * norm x ^ ?n" . from mult_right_le_imp_le[OF this xn0] have "norm x - 1 \ ml * (1 - 1 / (norm x ^ ?n))" by simp hence "norm x \ 1 + ml - ml / (norm x ^ ?n)" by (simp add: field_simps) also have "\ \ 1 + ml" using ml0 xn0 by auto finally show ?thesis unfolding ml_def . qed qed lemma div_le_div_ceiling: "x div y \ div_ceiling x y" unfolding div_ceiling_def Let_def by auto lemma div_ceiling: assumes q: "q \ 0" shows "(of_int x :: 'a :: floor_ceiling) / of_int q \ of_int (div_ceiling x q)" proof (cases "q dvd x") case True then obtain k where xqk: "x = q * k" unfolding dvd_def by auto hence id: "div_ceiling x q = k" unfolding div_ceiling_def Let_def using q by auto show ?thesis unfolding id unfolding xqk using q by simp next case False { assume "x div q * q = x" hence "x = q * (x div q)" by (simp add: ac_simps) hence "q dvd x" unfolding dvd_def by auto with False have False by simp } hence id: "div_ceiling x q = x div q + 1" unfolding div_ceiling_def Let_def using q by auto show ?thesis unfolding id by (metis floor_divide_of_int_eq le_less add1_zle_eq floor_less_iff) qed lemma max_list_non_empty_map: assumes hom: "\ x y. max (f x) (f y) = f (max x y)" shows "xs \ [] \ max_list_non_empty (map f xs) = f (max_list_non_empty xs)" by (induct xs rule: max_list_non_empty.induct, auto simp: hom) lemma root_bound: assumes "root_bound p = B" and deg: "degree p > 0" shows "ipoly p (x :: real) = 0 \ norm x \ of_rat B" "B \ 0" proof - let ?r = real_of_rat let ?i = real_of_int let ?p = "real_of_int_poly p" define n where "n = degree p" let ?lc = "coeff p n" let ?list = "map (\i. abs (coeff p i)) [0.. 0" by auto from p0 have alc0: "abs ?lc \ 0" unfolding n_def by auto from deg have mem: "abs (coeff p 0) \ set ?list" unfolding n_def by auto from max_list_non_empty[OF this, folded m_def] have m0: "m \ 0" by auto have "div_ceiling m (abs ?lc) \ 0" by (rule order_trans[OF _ div_le_div_ceiling[of m "abs ?lc"]], subst pos_imp_zdiv_nonneg_iff, insert p0 m0, auto simp: n_def) hence mup: "m_up \ 1" unfolding m_up_def by auto have "m_up \ 2 ^ (log_ceiling 2 m_up)" using mup log_ceiling_sound(1) by auto hence Cmup: "C \ of_int m_up" unfolding C_def by linarith with mup have C: "C \ 1" by auto from assms(1)[unfolded root_bound_def Let_def] have B: "C = of_rat B" unfolding C_def m_up_def n_def m_def by auto note dc = div_le_div_ceiling[of m "abs ?lc"] with C show "B \ 0" unfolding B by auto assume "ipoly p x = 0" hence rt: "poly ?p x = 0" by simp from root_imp_deg_nonzero[OF _ this] p0 have n0: "n \ 0" unfolding n_def by auto from cauchy_root_bound[OF rt] p0 have "norm x \ 1 + max_list_non_empty ?list' / ?i (abs ?lc)" by (simp add: n_def) also have "?list' = map ?i ?list" by simp also have "max_list_non_empty \ = ?i m" unfolding m_def by (rule max_list_non_empty_map, insert mem, auto) also have "1 + m / ?i (abs ?lc) \ ?i m_up" unfolding m_up_def using div_ceiling[OF alc0, of m] by auto also have "\ \ ?r C" using Cmup using of_rat_less_eq by force finally have "norm x \ ?r C" . thus "norm x \ ?r B" unfolding B by simp qed fun pairwise_disjoint :: "'a set list \ bool" where "pairwise_disjoint [] = True" | "pairwise_disjoint (x # xs) = ((x \ (\ y \ set xs. y) = {}) \ pairwise_disjoint xs)" lemma roots_of_2_irr: assumes pc: "poly_cond p" and deg: "degree p > 0" shows "real_of_2 ` set (roots_of_2_irr p) = {x. ipoly p x = 0}" (is ?one) "Ball (set (roots_of_2_irr p)) invariant_2" (is ?two) "distinct (map real_of_2 (roots_of_2_irr p))" (is ?three) proof - note d = roots_of_2_irr_def from poly_condD[OF pc] have mon: "lead_coeff p > 0" and irr: "irreducible p" by auto let ?norm = "real_alg_2'" have "?one \ ?two \ ?three" proof (cases "degree p = 1") case True define c where "c = coeff p 0" define d where "d = coeff p 1" from True have rr: "roots_of_2_irr p = [Rational (Rat.Fract (- c) (d))]" unfolding d d_def c_def by auto from degree1_coeffs[OF True] have p: "p = [:c,d:]" and d: "d \ 0" unfolding c_def d_def by auto have *: "real_of_int c + x * real_of_int d = 0 \ x = - (real_of_int c / real_of_int d)" for x using d by (simp add: field_simps) show ?thesis unfolding rr using d * unfolding p using of_rat_1[of "Rat.Fract (- c) (d)"] by (auto simp: Fract_of_int_quotient hom_distribs) next case False let ?r = real_of_rat let ?rp = "map_poly ?r" let ?rr = "set (roots_of_2_irr p)" define ri where "ri = root_info p" define cr where "cr = root_info.l_r ri" define bnds where "bnds = [(-root_bound p, root_bound p)]" define empty where "empty = (Nil :: real_alg_2 list)" have empty: "Ball (set empty) invariant_2 \ distinct (map real_of_2 empty)" unfolding empty_def by auto from mon have p: "p \ 0" by auto from root_info[OF irr deg] have ri: "root_info_cond ri p" unfolding ri_def . from False have rr: "roots_of_2_irr p = roots_of_2_main p ri cr bnds empty" unfolding d ri_def cr_def Let_def bnds_def empty_def by auto note root_bound = root_bound[OF refl deg] from root_bound(2) have bnds: "\ l r. (l,r) \ set bnds \ l \ r" unfolding bnds_def by auto have "ipoly p x = 0 \ ?r (- root_bound p) \ x \ x \ ?r (root_bound p)" for x using root_bound(1)[of x] by (auto simp: hom_distribs) hence rts: "{x. ipoly p x = 0} = real_of_2 ` set empty \ {x. \ l r. root_cond (p,l,r) x \ (l,r) \ set bnds}" unfolding empty_def bnds_def by (force simp: root_cond_def) define rts where "rts lr = Collect (root_cond (p,lr))" for lr have disj: "pairwise_disjoint (real_of_2 ` set empty # map rts bnds)" unfolding empty_def bnds_def by auto from deg False have deg1: "degree p > 1" by auto define delta where "delta = ipoly_root_delta p" note delta = ipoly_root_delta[OF p, folded delta_def] define rel' where "rel' = ({(x, y). 0 \ y \ delta_gt delta x y})^-1" define mm where "mm = (\bnds. mset (map (\ (l,r). ?r r - ?r l) bnds))" define rel where "rel = inv_image (mult1 rel') mm" have wf: "wf rel" unfolding rel_def rel'_def by (rule wf_inv_image[OF wf_mult1[OF SN_imp_wf[OF delta_gt_SN[OF delta(1)]]]]) let ?main = "roots_of_2_main p ri cr" have "real_of_2 ` set (?main bnds empty) = real_of_2 ` set empty \ {x. \l r. root_cond (p, l, r) x \ (l, r) \ set bnds} \ Ball (set (?main bnds empty)) invariant_2 \ distinct (map real_of_2 (?main bnds empty))" (is "?one' \ ?two' \ ?three'") using empty bnds disj proof (induct bnds arbitrary: empty rule: wf_induct[OF wf]) case (1 lrss rais) note rais = 1(2)[rule_format] note lrs = 1(3) note disj = 1(4) note IH = 1(1)[rule_format] note simp = roots_of_2_main.simps[of p ri cr lrss rais] show ?case proof (cases lrss) case Nil with rais show ?thesis unfolding simp by auto next case (Cons lr lrs) obtain l r where lr': "lr = (l,r)" by force { fix lr' assume lt: "\ l' r'. (l',r') \ set lr' \ l' \ r' \ delta_gt delta (?r r - ?r l) (?r r' - ?r l')" have l: "mm (lr' @ lrs) = mm lrs + mm lr'" unfolding mm_def by (auto simp: ac_simps) have r: "mm lrss = mm lrs + {# ?r r - ?r l #}" unfolding Cons lr' rel_def mm_def by auto have "(mm (lr' @ lrs), mm lrss) \ mult1 rel'" unfolding l r mult1_def proof (rule, unfold split, intro exI conjI, unfold add_mset_add_single[symmetric], rule refl, rule refl, intro allI impI) fix d assume "d \# mm lr'" then obtain l' r' where d: "d = ?r r' - ?r l'" and lr': "(l',r') \ set lr'" unfolding mm_def in_multiset_in_set by auto from lt[OF lr'] show "(d, ?r r - ?r l) \ rel'" unfolding d rel'_def by (auto simp: of_rat_less_eq) qed hence "(lr' @ lrs, lrss) \ rel" unfolding rel_def by auto } note rel = this from rel[of Nil] have easy_rel: "(lrs,lrss) \ rel" by auto define c where "c = cr l r" from simp Cons lr' have simp: "?main lrss rais = (if c = 0 then ?main lrs rais else if c = 1 then ?main lrs (real_alg_2' ri p l r # rais) else let m = (l + r) / 2 in ?main ((m, r) # (l, m) # lrs) rais)" unfolding c_def simp Cons lr' using real_alg_2''[OF False] by auto note lrs = lrs[unfolded Cons lr'] from lrs have lr: "l \ r" by auto from root_info_condD(1)[OF ri lr, folded cr_def] have c: "c = card {x. root_cond (p,l,r) x}" unfolding c_def by auto let ?rt = "\ lrs. {x. \l r. root_cond (p, l, r) x \ (l, r) \ set lrs}" have rts: "?rt lrss = ?rt lrs \ {x. root_cond (p,l,r) x}" (is "?rt1 = ?rt2 \ ?rt3") unfolding Cons lr' by auto show ?thesis proof (cases "c = 0") case True with simp have simp: "?main lrss rais = ?main lrs rais" by simp from disj have disj: "pairwise_disjoint (real_of_2 ` set rais # map rts lrs)" unfolding Cons by auto from finite_ipoly_roots[OF p] True[unfolded c] have empty: "?rt3 = {}" unfolding root_cond_def[abs_def] split by simp with rts have rts: "?rt1 = ?rt2" by auto show ?thesis unfolding simp rts by (rule IH[OF easy_rel rais lrs disj], auto) next case False show ?thesis proof (cases "c = 1") case True let ?rai = "real_alg_2' ri p l r" from True simp have simp: "?main lrss rais = ?main lrs (?rai # rais)" by auto from card_1_Collect_ex1[OF c[symmetric, unfolded True]] have ur: "unique_root (p,l,r)" . from real_alg_2'[OF ur pc ri] have rai: "invariant_2 ?rai" "real_of_2 ?rai = the_unique_root (p, l, r)" by auto with rais have rais: "\ x. x \ set (?rai # rais) \ invariant_2 x" and dist: "distinct (map real_of_2 rais)" by auto have rt3: "?rt3 = {real_of_2 ?rai}" using ur rai by (auto intro: the_unique_root_eqI theI') have "real_of_2 ` set (roots_of_2_main p ri cr lrs (?rai # rais)) = real_of_2 ` set (?rai # rais) \ ?rt2 \ Ball (set (roots_of_2_main p ri cr lrs (?rai # rais))) invariant_2 \ distinct (map real_of_2 (roots_of_2_main p ri cr lrs (?rai # rais)))" (is "?one \ ?two \ ?three") proof (rule IH[OF easy_rel, of "?rai # rais", OF conjI lrs]) show "Ball (set (real_alg_2' ri p l r # rais)) invariant_2" using rais by auto have "real_of_2 (real_alg_2' ri p l r) \ set (map real_of_2 rais)" using disj rt3 unfolding Cons lr' rts_def by auto thus "distinct (map real_of_2 (real_alg_2' ri p l r # rais))" using dist by auto show "pairwise_disjoint (real_of_2 ` set (real_alg_2' ri p l r # rais) # map rts lrs)" using disj rt3 unfolding Cons lr' rts_def by auto qed auto hence ?one ?two ?three by blast+ show ?thesis unfolding simp rts rt3 by (rule conjI[OF _ conjI[OF \?two\ \?three\]], unfold \?one\, auto) next case False let ?m = "(l+r)/2" let ?lrs = "[(?m,r),(l,?m)] @ lrs" from False \c \ 0\ have simp: "?main lrss rais = ?main ?lrs rais" unfolding simp by (auto simp: Let_def) from False \c \ 0\ have "c \ 2" by auto from delta(2)[OF this[unfolded c]] have delta: "delta \ ?r (r - l) / 4" by auto have lrs: "\ l r. (l,r) \ set ?lrs \ l \ r" using lr lrs by (fastforce simp: field_simps) have "?r ?m \ \" unfolding Rats_def by blast with poly_cond_degree_gt_1[OF pc deg1, of "?r ?m"] have disj1: "?r ?m \ rts lr" for lr unfolding rts_def root_cond_def by auto have disj2: "rts (?m, r) \ rts (l, ?m) = {}" using disj1[of "(l,?m)"] disj1[of "(?m,r)"] unfolding rts_def root_cond_def by auto have disj3: "(rts (l,?m) \ rts (?m,r)) = rts (l,r)" unfolding rts_def root_cond_def by (auto simp: hom_distribs) have disj4: "real_of_2 ` set rais \ rts (l,r) = {}" using disj unfolding Cons lr' by auto have disj: "pairwise_disjoint (real_of_2 ` set rais # map rts ([(?m, r), (l, ?m)] @ lrs))" using disj disj2 disj3 disj4 by (auto simp: Cons lr') have "(?lrs,lrss) \ rel" proof (rule rel, intro conjI) fix l' r' assume mem: "(l', r') \ set [(?m,r),(l,?m)]" from mem lr show "l' \ r'" by auto from mem have diff: "?r r' - ?r l' = (?r r - ?r l) / 2" by auto (metis eq_diff_eq minus_diff_eq mult_2_right of_rat_add of_rat_diff, metis of_rat_add of_rat_mult of_rat_numeral_eq) show "delta_gt delta (?r r - ?r l) (?r r' - ?r l')" unfolding diff delta_gt_def by (rule order.trans[OF delta], insert lr, auto simp: field_simps of_rat_diff of_rat_less_eq) qed note IH = IH[OF this, of rais, OF rais lrs disj] have "real_of_2 ` set (?main ?lrs rais) = real_of_2 ` set rais \ ?rt ?lrs \ Ball (set (?main ?lrs rais)) invariant_2 \ distinct (map real_of_2 (?main ?lrs rais))" (is "?one \ ?two") by (rule IH) hence ?one ?two by blast+ have cong: "\ a b c. b = c \ a \ b = a \ c" by auto have id: "?rt ?lrs = ?rt lrs \ ?rt [(?m,r),(l,?m)]" by auto show ?thesis unfolding rts simp \?one\ id proof (rule conjI[OF cong[OF cong] conjI]) have "\ x. root_cond (p,l,r) x = (root_cond (p,l,?m) x \ root_cond (p,?m,r) x)" unfolding root_cond_def by (auto simp:hom_distribs) hence id: "Collect (root_cond (p,l,r)) = {x. (root_cond (p,l,?m) x \ root_cond (p,?m,r) x)}" by auto show "?rt [(?m,r),(l,?m)] = Collect (root_cond (p,l,r))" unfolding id list.simps by blast show "\ a \ set (?main ?lrs rais). invariant_2 a" using \?two\ by auto show "distinct (map real_of_2 (?main ?lrs rais))" using \?two\ by auto qed qed qed qed qed hence idd: "?one'" and cond: ?two' ?three' by blast+ define res where "res = roots_of_2_main p ri cr bnds empty" have e: "set empty = {}" unfolding empty_def by auto from idd[folded res_def] e have idd: "real_of_2 ` set res = {} \ {x. \l r. root_cond (p, l, r) x \ (l, r) \ set bnds}" by auto show ?thesis unfolding rr unfolding rts id e norm_def using cond unfolding res_def[symmetric] image_empty e idd[symmetric] by auto qed thus ?one ?two ?three by blast+ qed definition roots_of_2 :: "int poly \ real_alg_2 list" where "roots_of_2 p = concat (map roots_of_2_irr (factors_of_int_poly p))" lemma roots_of_2: shows "p \ 0 \ real_of_2 ` set (roots_of_2 p) = {x. ipoly p x = 0}" "Ball (set (roots_of_2 p)) invariant_2" "distinct (map real_of_2 (roots_of_2 p))" proof - let ?rr = "roots_of_2 p" note d = roots_of_2_def note frp1 = factors_of_int_poly { fix q r assume "q \ set ?rr" then obtain s where s: "s \ set (factors_of_int_poly p)" and q: "q \ set (roots_of_2_irr s)" unfolding d by auto from frp1(1)[OF refl s] have "poly_cond s" "degree s > 0" by (auto simp: poly_cond_def) from roots_of_2_irr[OF this] q have "invariant_2 q" by auto } thus "Ball (set ?rr) invariant_2" by auto { assume p: "p \ 0" have "real_of_2 ` set ?rr = (\ ((\ p. real_of_2 ` set (roots_of_2_irr p)) ` (set (factors_of_int_poly p))))" (is "_ = ?rrr") unfolding d set_concat set_map by auto also have "\ = {x. ipoly p x = 0}" proof - { fix x assume "x \ ?rrr" then obtain q s where s: "s \ set (factors_of_int_poly p)" and q: "q \ set (roots_of_2_irr s)" and x: "x = real_of_2 q" by auto from frp1(1)[OF refl s] have s0: "s \ 0" and pt: "poly_cond s" "degree s > 0" by (auto simp: poly_cond_def) from roots_of_2_irr[OF pt] q have rt: "ipoly s x = 0" unfolding x by auto from frp1(2)[OF refl p, of x] rt s have rt: "ipoly p x = 0" by auto } moreover { fix x :: real assume rt: "ipoly p x = 0" from rt frp1(2)[OF refl p, of x] obtain s where s: "s \ set (factors_of_int_poly p)" and rt: "ipoly s x = 0" by auto from frp1(1)[OF refl s] have s0: "s \ 0" and ty: "poly_cond s" "degree s > 0" by (auto simp: poly_cond_def) from roots_of_2_irr(1)[OF ty] rt obtain q where q: "q \ set (roots_of_2_irr s)" and x: "x = real_of_2 q" by blast have "x \ ?rrr" unfolding x using q s by auto } ultimately show ?thesis by auto qed finally show "real_of_2 ` set ?rr = {x. ipoly p x = 0}" by auto } show "distinct (map real_of_2 (roots_of_2 p))" proof (cases "p = 0") case True from factors_of_int_poly_const[of 0] True show ?thesis unfolding roots_of_2_def by auto next case p: False note frp1 = frp1[OF refl] let ?fp = "factors_of_int_poly p" let ?cc = "concat (map roots_of_2_irr ?fp)" show ?thesis unfolding roots_of_2_def distinct_conv_nth length_map proof (intro allI impI notI) fix i j assume ij: "i < length ?cc" "j < length ?cc" "i \ j" and id: "map real_of_2 ?cc ! i = map real_of_2 ?cc ! j" from ij id have id: "real_of_2 (?cc ! i) = real_of_2 (?cc ! j)" by auto from nth_concat_diff[OF ij, unfolded length_map] obtain j1 k1 j2 k2 where *: "(j1,k1) \ (j2,k2)" "j1 < length ?fp" "j2 < length ?fp" and "k1 < length (map roots_of_2_irr ?fp ! j1)" "k2 < length (map roots_of_2_irr ?fp ! j2)" "?cc ! i = map roots_of_2_irr ?fp ! j1 ! k1" "?cc ! j = map roots_of_2_irr ?fp ! j2 ! k2" by blast hence **: "k1 < length (roots_of_2_irr (?fp ! j1))" "k2 < length (roots_of_2_irr (?fp ! j2))" "?cc ! i = roots_of_2_irr (?fp ! j1) ! k1" "?cc ! j = roots_of_2_irr (?fp ! j2) ! k2" by auto from * have mem: "?fp ! j1 \ set ?fp" "?fp ! j2 \ set ?fp" by auto from frp1(1)[OF mem(1)] frp1(1)[OF mem(2)] have pc1: "poly_cond (?fp ! j1)" "degree (?fp ! j1) > 0" and pc10: "?fp ! j1 \ 0" and pc2: "poly_cond (?fp ! j2)" "degree (?fp ! j2) > 0" by (auto simp: poly_cond_def) show False proof (cases "j1 = j2") case True with * have neq: "k1 \ k2" by auto from **[unfolded True] id * have "map real_of_2 (roots_of_2_irr (?fp ! j2)) ! k1 = real_of_2 (?cc ! j)" "map real_of_2 (roots_of_2_irr (?fp ! j2)) ! k1 = real_of_2 (?cc ! j)" by auto hence "\ distinct (map real_of_2 (roots_of_2_irr (?fp ! j2)))" unfolding distinct_conv_nth using * ** True by auto with roots_of_2_irr(3)[OF pc2] show False by auto next case neq: False with frp1(4)[of p] * have neq: "?fp ! j1 \ ?fp ! j2" unfolding distinct_conv_nth by auto let ?x = "real_of_2 (?cc ! i)" define x where "x = ?x" from ** have "x \ real_of_2 ` set (roots_of_2_irr (?fp ! j1))" unfolding x_def by auto with roots_of_2_irr(1)[OF pc1] have x1: "ipoly (?fp ! j1) x = 0" by auto from ** id have "x \ real_of_2 ` set (roots_of_2_irr (?fp ! j2))" unfolding x_def by (metis image_eqI nth_mem) with roots_of_2_irr(1)[OF pc2] have x2: "ipoly (?fp ! j2) x = 0" by auto have "ipoly p x = 0" using x1 mem unfolding roots_of_2_def by (metis frp1(2) p) from frp1(3)[OF p this] x1 x2 neq mem show False by blast qed qed qed qed lift_definition roots_of_3 :: "int poly \ real_alg_3 list" is roots_of_2 by (insert roots_of_2, auto simp: list_all_iff) lemma roots_of_3: shows "p \ 0 \ real_of_3 ` set (roots_of_3 p) = {x. ipoly p x = 0}" "distinct (map real_of_3 (roots_of_3 p))" proof - show "p \ 0 \ real_of_3 ` set (roots_of_3 p) = {x. ipoly p x = 0}" by (transfer; intro roots_of_2, auto) show "distinct (map real_of_3 (roots_of_3 p))" by (transfer; insert roots_of_2, auto) qed lift_definition roots_of_real_alg :: "int poly \ real_alg list" is roots_of_3 . lemma roots_of_real_alg: "p \ 0 \ real_of ` set (roots_of_real_alg p) = {x. ipoly p x = 0}" "distinct (map real_of (roots_of_real_alg p))" proof - show "p \ 0 \ real_of ` set (roots_of_real_alg p) = {x. ipoly p x = 0}" by (transfer', insert roots_of_3, auto) show "distinct (map real_of (roots_of_real_alg p))" by (transfer, insert roots_of_3(2), auto) qed text \It follows an implementation for @{const roots_of_3}, since the current definition does not provide a code equation.\ context begin private typedef real_alg_2_list = "{xs. Ball (set xs) invariant_2}" by (intro exI[of _ Nil], auto) setup_lifting type_definition_real_alg_2_list private lift_definition roots_of_2_list :: "int poly \ real_alg_2_list" is roots_of_2 by (insert roots_of_2, auto) private lift_definition real_alg_2_list_nil :: "real_alg_2_list \ bool" is "\ xs. case xs of Nil \ True | _ \ False" . private fun real_alg_2_list_hd_intern :: "real_alg_2 list \ real_alg_2" where "real_alg_2_list_hd_intern (Cons x xs) = x" | "real_alg_2_list_hd_intern Nil = of_rat_2 0" private lift_definition real_alg_2_list_hd :: "real_alg_2_list \ real_alg_3" is real_alg_2_list_hd_intern proof (goal_cases) case (1 xs) thus ?case using of_rat_2[of 0] by (cases xs, auto) qed private lift_definition real_alg_2_list_tl :: "real_alg_2_list \ real_alg_2_list" is tl proof (goal_cases) case (1 xs) thus ?case by (cases xs, auto) qed private lift_definition real_alg_2_list_length :: "real_alg_2_list \ nat" is length . private lemma real_alg_2_list_length[simp]: "\ real_alg_2_list_nil xs \ real_alg_2_list_length (real_alg_2_list_tl xs) < real_alg_2_list_length xs" by (transfer, auto split: list.splits) private function real_alg_2_list_convert :: "real_alg_2_list \ real_alg_3 list" where "real_alg_2_list_convert xs = (if real_alg_2_list_nil xs then [] else real_alg_2_list_hd xs # real_alg_2_list_convert (real_alg_2_list_tl xs))" by pat_completeness auto termination by (relation "measure real_alg_2_list_length", auto) private definition roots_of_3_impl :: "int poly \ real_alg_3 list" where "roots_of_3_impl p = real_alg_2_list_convert (roots_of_2_list p)" private lift_definition real_alg_2_list_convert_id :: "real_alg_2_list \ real_alg_3 list" is id by (auto simp: list_all_iff) lemma real_alg_2_list_convert: "real_alg_2_list_convert xs = real_alg_2_list_convert_id xs" proof (induct xs rule: wf_induct[OF wf_measure[of real_alg_2_list_length], rule_format]) case (1 xs) show ?case proof (cases "real_alg_2_list_nil xs") case True hence "real_alg_2_list_convert xs = []" by auto also have "[] = real_alg_2_list_convert_id xs" using True by (transfer', auto split: list.splits) finally show ?thesis . next case False hence "real_alg_2_list_convert xs = real_alg_2_list_hd xs # real_alg_2_list_convert (real_alg_2_list_tl xs)" by simp also have "real_alg_2_list_convert (real_alg_2_list_tl xs) = real_alg_2_list_convert_id (real_alg_2_list_tl xs)" by (rule 1, insert False, simp) also have "real_alg_2_list_hd xs # \ = real_alg_2_list_convert_id xs" using False by (transfer', auto split: list.splits) finally show ?thesis . qed qed lemma roots_of_3_code[code]: "roots_of_3 p = roots_of_3_impl p" unfolding roots_of_3_impl_def real_alg_2_list_convert by (transfer, simp) end definition real_roots_of_int_poly :: "int poly \ real list" where "real_roots_of_int_poly p = map real_of (roots_of_real_alg p)" definition real_roots_of_rat_poly :: "rat poly \ real list" where "real_roots_of_rat_poly p = map real_of (roots_of_real_alg (snd (rat_to_int_poly p)))" abbreviation rpoly :: "rat poly \ 'a :: field_char_0 \ 'a" where "rpoly f \ poly (map_poly of_rat f)" lemma real_roots_of_int_poly: "p \ 0 \ set (real_roots_of_int_poly p) = {x. ipoly p x = 0}" "distinct (real_roots_of_int_poly p)" unfolding real_roots_of_int_poly_def using roots_of_real_alg[of p] by auto lemma real_roots_of_rat_poly: "p \ 0 \ set (real_roots_of_rat_poly p) = {x. rpoly p x = 0}" "distinct (real_roots_of_rat_poly p)" proof - obtain c q where cq: "rat_to_int_poly p = (c,q)" by force from rat_to_int_poly[OF this] have pq: "p = smult (inverse (of_int c)) (of_int_poly q)" and c: "c \ 0" by auto have id: "{x. rpoly p x = (0 :: real)} = {x. ipoly q x = 0}" unfolding pq by (simp add: c of_rat_of_int_poly hom_distribs) show "distinct (real_roots_of_rat_poly p)" unfolding real_roots_of_rat_poly_def cq snd_conv using roots_of_real_alg(2)[of q] . assume "p \ 0" with pq c have q: "q \ 0" by auto show "set (real_roots_of_rat_poly p) = {x. rpoly p x = 0}" unfolding id unfolding real_roots_of_rat_poly_def cq snd_conv using roots_of_real_alg(1)[OF q] by auto qed -text \The upcoming functions no longer demand an integer or rational polynomial as input.\ - -definition roots_of_real_main :: "real poly \ real list" where - "roots_of_real_main p \ let n = degree p in - if n = 0 then [] else if n = 1 then [roots1 p] else if n = 2 then rroots2 p - else (real_roots_of_rat_poly (map_poly to_rat p))" - -definition roots_of_real_poly :: "real poly \ real list option" where - "roots_of_real_poly p \ let (c,pis) = yun_factorization gcd p in - if (c \ 0 \ (\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then - Some (concat (map (roots_of_real_main o fst) pis)) else None" - -lemma roots_of_real_main: assumes p: "p \ 0" and deg: "degree p \ 2 \ set (coeffs p) \ \" - shows "set (roots_of_real_main p) = {x. poly p x = 0}" (is "?l = ?r") -proof - - note d = roots_of_real_main_def Let_def - show ?thesis - proof (cases "degree p = 0") - case True - hence "?l = {}" unfolding d by auto - with roots0[OF p True] show ?thesis by auto - next - case False note 0 = this - show ?thesis - proof (cases "degree p = 1") - case True - hence "?l = {roots1 p}" unfolding d by auto - with roots1[OF True] show ?thesis by auto - next - case False note 1 = this - show ?thesis - proof (cases "degree p = 2") - case True - hence "?l = set (rroots2 p)" unfolding d by auto - with rroots2[OF True] show ?thesis by auto - next - case False note 2 = this - let ?q = "map_poly to_rat p" - from 0 1 2 have l: "?l = set (real_roots_of_rat_poly ?q)" unfolding d by auto - from deg 0 1 2 have rat: "set (coeffs p) \ \" by auto - have "p = map_poly (of_rat o to_rat) p" - by (rule sym, rule map_poly_idI, insert rat, auto) - also have "\ = real_of_rat_poly ?q" - by (subst map_poly_map_poly, auto simp: to_rat) - finally have id: "{x. poly p x = 0} = {x. poly (real_of_rat_poly ?q) x = 0}" and q: "?q \ 0" - using p by auto - from real_roots_of_rat_poly(1)[OF q, folded id l] show ?thesis by simp - qed - qed - qed -qed - -lemma roots_of_real_poly: assumes rt: "roots_of_real_poly p = Some xs" - shows "set xs = {x. poly p x = 0}" -proof - - obtain c pis where yun: "yun_factorization gcd p = (c,pis)" by force - from rt[unfolded roots_of_real_poly_def yun split Let_def] - have c: "c \ 0" and pis: "\ p i. (p, i)\set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" - and xs: "xs = concat (map (roots_of_real_main \ fst) pis)" - by (auto split: if_splits) - note yun = square_free_factorizationD(1,2,4)[OF yun_factorization(1)[OF yun]] - from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . - have "{x. poly p x = 0} = {x. poly (\(a, i)\set pis. a ^ Suc i) x = 0}" - unfolding p using c by auto - also have "\ = \ ((\ p. {x. poly p x = 0}) ` fst ` set pis)" (is "_ = ?r") - by (subst poly_prod_0, force+) - finally have r: "{x. poly p x = 0} = ?r" . - { - fix p i - assume p: "(p,i) \ set pis" - have "set (roots_of_real_main p) = {x. poly p x = 0}" - by (rule roots_of_real_main, insert yun(2)[OF p] pis[OF p], auto) - } note main = this - have "set xs = \ ((\ (p, i). set (roots_of_real_main p)) ` set pis)" unfolding xs o_def - by auto - also have "\ = ?r" using main by auto - finally show ?thesis unfolding r by simp -qed - end diff --git a/thys/Cubic_Quartic_Equations/Complex_Roots.thy b/thys/Cubic_Quartic_Equations/Complex_Roots.thy --- a/thys/Cubic_Quartic_Equations/Complex_Roots.thy +++ b/thys/Cubic_Quartic_Equations/Complex_Roots.thy @@ -1,454 +1,516 @@ section \$n$-th roots of complex numbers\ theory Complex_Roots imports Complex_Geometry.More_Complex - Min_Int_Poly_Complex_Impl + Algebraic_Numbers.Complex_Algebraic_Numbers + Factor_Algebraic_Polynomial.Roots_via_IA "HOL-Library.Product_Lexorder" begin subsection \An algorithm to compute all complex roots of (algebraic) complex numbers\ -text \TODO: The filter instruction might be tuned by using interval arithmetic instead.\ definition all_croots :: "nat \ complex \ complex list" where "all_croots n x = (if n = 0 then [] else if algebraic x then (let p = min_int_poly x; q = poly_nth_root n p; xs = complex_roots_of_int_poly q in filter (\ y. y^n = x) xs) else (SOME ys. set ys = {y. y^n = x}))" -lemma all_croots_code[code]: - "all_croots n x = (if n = 0 then [] else - if algebraic x then - (let p = min_int_poly x; - q = poly_nth_root n p; - xs = complex_roots_of_int_poly q - in filter (\ y. y^n = x) xs) - else Code.abort (STR ''all_croots invoked on non-algebraic number'') (\ _. all_croots n x))" - by (auto simp: all_croots_def) - lemma all_croots: assumes n0: "n \ 0" shows "set (all_croots n x) = {y. y^n = x}" proof (cases "algebraic x") case True hence id: "(if n = 0 then y else if algebraic x then z else u) = z" for y z u :: "complex list" using n0 by auto define p where "p = poly_nth_root n (min_int_poly x)" show ?thesis unfolding Let_def p_def[symmetric] all_croots_def id proof (standard, force, standard, simp) fix y assume y: "y ^n = x" have "min_int_poly x represents x" using True by auto from represents_nth_root[OF n0 y this] have "p represents y" unfolding p_def by auto thus "y \ set (complex_roots_of_int_poly p)" by (subst complex_roots_of_int_poly, auto) qed next case False hence id: "(if n = 0 then y else if algebraic x then z else u) = u" for y z u :: "complex list" using n0 by auto show ?thesis unfolding Let_def all_croots_def id by (rule someI_ex, rule finite_list, insert n0, blast) qed +text \TODO: One might change @{const complex_roots_of_int_poly} to @{const complex_roots_of_int_poly3} + in order to avoid an unnecessary factorization of an integer polynomial. However, then + this change already needs to be performed within the definition of @{const all_croots}.\ +lift_definition all_croots_part1 :: "nat \ complex \ complex genuine_roots_aux" is + "\ n x. if n = 0 \ x = 0 \ \ algebraic x then (1,[],0, filter_fun_complex 1) + else let p = min_int_poly x; + q = poly_nth_root n p; + zeros = complex_roots_of_int_poly q; + r = Polynomial.monom 1 n - [:x:] + in (r,zeros, n, filter_fun_complex r)" + subgoal for n x + proof (cases "n = 0 \ x = 0 \ \ algebraic x") + case True + thus ?thesis by (simp add: filter_fun_complex) + next + case False + hence *: "algebraic x" "n \ 0" "x \ 0" by auto + { + fix z + assume zn: "z^n = x" + from *(1) have repr: "min_int_poly x represents x" by auto + from represents_nth_root[OF *(2) zn repr] + have "poly_nth_root n (min_int_poly x) represents z" . + } + moreover have "card {z. z ^ n = x} = n" + by (rule card_nth_roots) (use * in auto) + ultimately show ?thesis using * + by (auto simp: Let_def complex_roots_of_int_poly filter_fun_complex poly_monom) + qed + done + +lemma all_croots_code[code]: + "all_croots n x = (if n = 0 then [] else if x = 0 then [0] + else if algebraic x then genuine_roots_impl (all_croots_part1 n x) + else Code.abort (STR ''all_croots invoked on non-algebraic number'') (\ _. all_croots n x))" +proof (cases "n = 0") + case True + thus ?thesis unfolding all_croots_def by simp +next + case n: False + show ?thesis + proof (cases "x = 0") + case x: False + show ?thesis + proof (cases "algebraic x") + case False + with n x show ?thesis by simp + next + case True + define t where "t = ?thesis" + have "t \ filter (\y. y ^ n = x) + (complex_roots_of_int_poly (poly_nth_root n (min_int_poly x))) + = genuine_roots_impl (all_croots_part1 n x)" + unfolding t_def + by (subst all_croots_def[of n x], unfold Let_def, insert n x True, auto) + also have "\" using n x True unfolding genuine_roots_impl_def + by (transfer, simp add: Let_def genuine_roots_def poly_monom) + finally show ?thesis unfolding t_def by simp + qed + next + case x: True + have "set (all_croots n 0) = {0}" unfolding all_croots[OF n] using n by simp + moreover have "distinct (all_croots n 0)" unfolding all_croots_def using n + by (auto intro!: distinct_filter complex_roots_of_int_poly) + ultimately have "all_croots n 0 = [0]" + by (smt (verit, del_insts) distinct.simps(2) distinct_singleton insert_ident list.set_cases list.set_intros(1) list.simps(15) mem_Collect_eq set_empty singleton_conv) + moreover have "?thesis \ all_croots n 0 = [0]" using n x by simp + ultimately show ?thesis by auto + qed +qed + + subsection \A definition of \emph{the} complex root of a complex number\ text \While the definition of the complex root is quite natural and easy, the main task is a criterion to determine which of all possible roots of a complex number is the chosen one.\ definition croot :: "nat \ complex \ complex" where "croot n x = (rcis (root n (cmod x)) (Arg x / of_nat n))" lemma croot_0[simp]: "croot n 0 = 0" "croot 0 x = 0" unfolding croot_def by auto lemma croot_power: assumes n: "n \ 0" shows "(croot n x) ^ n = x" unfolding croot_def DeMoivre2 by (subst real_root_pow_pos2, insert n, auto simp: rcis_cmod_Arg) lemma Arg_of_real: "Arg (of_real x) = (if x < 0 then pi else 0)" proof (cases "x = 0") case False hence "x < 0 \ x > 0" by auto thus ?thesis by (intro cis_Arg_unique, auto simp: complex_sgn_def scaleR_complex.ctr complex_eq_iff) qed (auto simp: Arg_def) lemma Arg_rcis_cis[simp]: assumes "x > 0" shows "Arg (rcis x y) = Arg (cis y)" using assms unfolding rcis_def by simp lemma cis_Arg_1[simp]: "cis (Arg 1) = 1" using Arg_of_real[of 1] by simp lemma cis_Arg_power[simp]: assumes "x \ 0" shows "cis (Arg (x ^ n)) = cis (Arg x * real n)" proof (induct n) case (Suc n) show ?case unfolding power.simps proof (subst cis_arg_mult) show "cis (Arg x + Arg (x ^ n)) = cis (Arg x * real (Suc n))" unfolding mult.commute[of "Arg x"] DeMoivre[symmetric] unfolding power.simps using Suc by (metis DeMoivre cis_mult mult.commute) show "x * x ^ n \ 0" using assms by auto qed qed simp lemma Arg_croot[simp]: "Arg (croot n x) = Arg x / real n" proof (cases "n = 0 \ x = 0") case True thus ?thesis by (auto simp: Arg_def) next case False hence n: "n \ 0" and x: "x \ 0" by auto let ?root = "croot n x" from n have n1: "real n \ 1" "real n > 0" "real n \ 0" by auto have bounded: "- pi < Arg x / real n \ Arg x / real n \ pi" proof (cases "Arg x < 0") case True from Arg_bounded[of x] have "- pi < Arg x" by auto also have "\ \ Arg x / real n" using n1 True by (smt (z3) div_by_1 divide_minus_left frac_le) finally have one: "- pi < Arg x / real n" . have "Arg x / real n \ 0" using True n1 by (smt (verit) divide_less_0_iff) also have "\ \ pi" by simp finally show ?thesis using one by auto next case False hence ax: "Arg x \ 0" by auto have "Arg x / real n \ Arg x" using n1 ax by (smt (verit) div_by_1 frac_le) also have "\ \ pi" using Arg_bounded[of x] by simp finally have one: "Arg x / real n \ pi" . have "-pi < 0" by simp also have "\ \ Arg x / real n" using ax n1 by simp finally show ?thesis using one by auto qed have "Arg ?root = Arg (cis (Arg x / real n))" unfolding croot_def using x n by simp also have "\ = Arg x / real n" by (rule cis_Arg_unique, force, insert bounded, auto) finally show ?thesis . qed lemma cos_abs[simp]: "cos (abs x :: real) = cos x" proof (cases "x < 0") case True hence abs: "abs x = - x" by simp show ?thesis unfolding abs by simp qed simp lemma cos_mono_le: assumes "abs x \ pi" and "abs y \ pi" shows "cos x \ cos y \ abs y \ abs x" proof - have "cos x \ cos y \ cos (abs x) \ cos (abs y)" by simp also have "\ \ abs y \ abs x" by (subst cos_mono_le_eq, insert assms, auto) finally show ?thesis . qed lemma abs_add_2_mult_bound: fixes x :: "'a :: linordered_idom" assumes xy: "\x\ \ y" shows "\x\ \ \x + 2 * of_int i * y\" proof (cases "i = 0") case i: False let ?oi = "of_int :: int \ 'a" from xy have y: "y \ 0" by auto consider (pp) "x \ 0" "i \ 0" | (nn) "x \ 0" "i \ 0" | (pn) "x \ 0" "i \ 0" | (np) "x \ 0" "i \ 0" by linarith thus ?thesis proof cases case pp thus ?thesis using y by simp next case nn have "x \ x + 2 * ?oi i * y" using nn y by (simp add: mult_nonneg_nonpos2) with nn show ?thesis by linarith next case pn with i have "0 \ x" "i < 0" by auto define j where "j = nat (-i) - 1" define z where "z = x - 2 * y" define u where "u = 2 * ?oi (nat j) * y" have u: "u \ 0" unfolding u_def using y by auto have i: "i = - int (Suc j)" using \i < 0\ unfolding j_def by simp have id: "x + 2 * ?oi i * y = z - u" unfolding i z_def u_def by (simp add: field_simps) have z: "z \ 0" "abs z \ x" using xy y pn(1) unfolding z_def by auto show ?thesis unfolding id using pn(1) z u by simp next case np with i have "0 \ x" "i > 0" by auto define j where "j = nat i - 1" have i: "i = int (Suc j)" using \i > 0\ unfolding j_def by simp define u where "u = 2 * ?oi (nat j) * y" have u: "u \ 0" unfolding u_def using y by auto define z where "z = - x - 2 * y" have id: "x + 2 * ?oi i * y = - z + u" unfolding i z_def u_def by (simp add: field_simps) have z: "z \ 0" "abs z \ - x" using xy y np(1) unfolding z_def by auto show ?thesis unfolding id using np(1) z u by simp qed qed simp lemma abs_eq_add_2_mult: fixes y :: "'a :: linordered_idom" assumes abs_id: "\x\ = \x + 2 * of_int i * y\" and xy: "- y < x" "x \ y" and i: "i \ 0" shows "x = y \ i = -1" proof - let ?oi = "of_int :: int \ 'a" from xy have y: "y > 0" by auto consider (pp) "x \ 0" "i \ 0" | (nn) "x < 0" "i \ 0" | (pn) "x \ 0" "i \ 0" | (np) "x < 0" "i \ 0" by linarith hence "?thesis \ x = ?oi (- i) * y" proof cases case pp thus ?thesis using y abs_id xy i by simp next case nn hence "\x + 2 * ?oi i * y\ = - (x + 2 * ?oi i * y)" using y nn by (intro abs_of_nonpos add_nonpos_nonpos, force, simp, intro mult_nonneg_nonpos, auto) thus ?thesis using y abs_id xy i nn by auto next case pn with i have "0 \ x" "i < 0" by auto define j where "j = nat (-i) - 1" define z where "z = x - 2 * y" define u where "u = 2 * ?oi (nat j) * y" have u: "u \ 0" unfolding u_def using y by auto have i: "i = - int (Suc j)" using \i < 0\ unfolding j_def by simp have id: "x + 2 * ?oi i * y = z - u" unfolding i z_def u_def by (simp add: field_simps) have z: "z \ 0" "abs z \ x" using xy y pn(1) unfolding z_def by auto from abs_id[unfolded id] have "z - u = -x " using z u pn by auto from this[folded id] have "x = of_int (-i) * y" by auto thus ?thesis by auto next case np with i have "0 \ x" "i > 0" by auto define j where "j = nat i - 1" have i: "i = int (Suc j)" using \i > 0\ unfolding j_def by simp define u where "u = 2 * ?oi (nat j) * y" have u: "u \ 0" unfolding u_def using y by auto define z where "z = - x - 2 * y" have id: "x + 2 * ?oi i * y = - z + u" unfolding i z_def u_def by (simp add: field_simps) have z: "z \ 0" using xy y np(1) unfolding z_def by auto from abs_id[unfolded id] have "- z + u = - x" using u z np by auto from this[folded id] have "x = of_int (- i) * y" by auto thus ?thesis by auto qed thus ?thesis proof assume "x = ?oi (- i) * y" with xy i y show ?thesis by (smt (verit, ccfv_SIG) less_le minus_less_iff mult_le_cancel_right2 mult_minus1_right mult_minus_left mult_of_int_commute of_int_hom.hom_one of_int_le_1_iff of_int_minus) qed qed text \This is the core lemma. It tells us that @{const croot} will choose the principal root, i.e. the root with largest real part and if there are two roots with identical real part, then the largest imaginary part. This criterion will be crucial for implementing @{const croot}.\ lemma croot_principal: assumes n: "n \ 0" and y: "y ^ n = x" and neq: "y \ croot n x" shows "Re y < Re (croot n x) \ Re y = Re (croot n x) \ Im y < Im (croot n x)" proof (cases "x = 0") case True with neq y have False by auto thus ?thesis .. next case x: False let ?root = "croot n x" from n have n1: "real n \ 1" "real n > 0" "real n \ 0" by auto from x y n have y0: "y \ 0" by auto from croot_power[OF n, of x] y have id: "?root ^ n = y ^ n" by simp hence "cmod (?root ^ n) = cmod (y ^ n)" by simp hence norm_eq: "cmod ?root = cmod y" using n unfolding norm_power by (meson gr_zeroI norm_ge_zero power_eq_imp_eq_base) have "cis (Arg y * real n) = cis (Arg (y^n))" by (subst cis_Arg_power[OF y0], simp) also have "\ = cis (Arg x)" using y by simp finally have ciseq: "cis (Arg y * real n) = cis (Arg x)" by simp from cis_eq[OF ciseq] obtain i where "Arg y * real n - Arg x = 2 * real_of_int i * pi" by auto hence "Arg y * real n = Arg x + 2 * real_of_int i * pi" by auto from arg_cong[OF this, of "\ x. x / real n"] n1 have Argy: "Arg y = Arg ?root + 2 * real_of_int i * pi / real n" by (auto simp: field_simps) have i0: "i \ 0" proof assume "i = 0" hence "Arg y = Arg ?root" unfolding Argy by simp with norm_eq have "?root = y" by (metis rcis_cmod_Arg) with neq show False by simp qed from y0 have cy0: "cmod y > 0" by auto from Arg_bounded[of x] have abs_pi: "abs (Arg x) \ pi" by auto have "Re y \ Re ?root \ Re y / cmod y \ Re ?root / cmod y" using cy0 unfolding divide_le_cancel by simp also have cosy: "Re y / cmod y = cos (Arg y)" unfolding cos_arg[OF y0] .. also have cosrt: "Re ?root / cmod y = cos (Arg ?root)" unfolding norm_eq[symmetric] by (subst cos_arg, insert norm_eq cy0, auto) also have "cos (Arg y) \ cos (Arg ?root) \ abs (Arg ?root) \ abs (Arg y)" by (rule cos_mono_le, insert Arg_bounded[of y] Arg_bounded[of ?root], auto) also have "\ \ abs (Arg ?root) * real n \ abs (Arg y) * real n" unfolding mult_le_cancel_right using n1 by simp also have "\ \ abs (Arg x) \ \Arg x + 2 * real_of_int i * pi\" unfolding Argy using n1 by (simp add: field_simps) also have "\" using abs_pi by (rule abs_add_2_mult_bound) finally have le: "Re y \ Re (croot n x)" . show ?thesis proof (cases "Re y = Re (croot n x)") case False with le show ?thesis by auto next case True hence "Re y / cmod y = Re ?root / cmod y" by simp hence "cos (Arg y) = cos (Arg ?root)" unfolding cosy cosrt . hence "cos (abs (Arg y)) = cos (abs (Arg ?root))" unfolding cos_abs . from cos_inj_pi[OF _ _ _ _ this] have "abs (Arg y) = abs (Arg ?root)" using Arg_bounded[of y] Arg_bounded[of ?root] by auto hence "abs (Arg y) * real n = abs (Arg ?root) * real n" by simp hence "abs (Arg x) = \Arg x + 2 * real_of_int i * pi\" unfolding Argy using n1 by (simp add: field_simps) from abs_eq_add_2_mult[OF this _ _ \i \ 0\] Arg_bounded[of x] have Argx: "Arg x = pi" and i: "i = -1" by auto have Argy: "Arg y = -pi / real n" unfolding Argy Arg_croot i Argx by simp have "Im ?root > Im y \ Im ?root / cmod ?root > Im y / cmod y" unfolding norm_eq using cy0 by (meson divide_less_cancel divide_strict_right_mono) also have "\ \ sin (Arg ?root) > sin (Arg y)" by (subst (1 2) sin_arg, insert y0 norm_eq, auto) also have "\ \ sin (- pi / real n) < sin (pi / real n)" unfolding Argy Arg_croot Argx by simp also have \ proof - have "sin (- pi / real n) < 0" using n1 by (smt (verit) Arg_bounded Argy divide_neg_pos sin_gt_zero sin_minus) also have "\ < sin (pi / real n)" using n1 calculation by fastforce finally show ?thesis . qed finally show ?thesis using le by auto qed qed lemma croot_unique: assumes n: "n \ 0" and y: "y ^ n = x" and y_max_Re_Im: "\ z. z ^ n = x \ Re z < Re y \ Re z = Re y \ Im z \ Im y" shows "croot n x = y" proof (rule ccontr) assume "croot n x \ y" from croot_principal[OF n y this[symmetric]] have "Re y < Re (croot n x) \ Re y = Re (croot n x) \ Im y < Im (croot n x)" . with y_max_Re_Im[OF croot_power[OF n]] show False by auto qed lemma csqrt_is_croot_2: "csqrt = croot 2" proof fix x show "csqrt x = croot 2 x" proof (rule sym, rule croot_unique, force, force) let ?p = "[:-x,0,1:]" let ?cx = "csqrt x" have p: "?p = [:?cx,1:] * [:-?cx,1:]" by (simp add: power2_eq_square[symmetric]) fix y assume "y^2 = x" hence "True \ poly ?p y = 0" by (auto simp: power2_eq_square) also have "\ \ y = - ?cx \ y = ?cx" unfolding p poly_mult mult_eq_0_iff poly_root_factor by auto finally have "y = - ?cx \ y = ?cx" by simp thus "Re y < Re ?cx \ Re y = Re ?cx \ Im y \ Im ?cx" proof assume y: "y = - ?cx" show ?thesis proof (cases "Re ?cx = 0") case False with csqrt_principal[of x] have "Re ?cx > 0" by simp thus ?thesis unfolding y by simp next case True with csqrt_principal[of x] have "Im ?cx \ 0" by simp thus ?thesis unfolding y using True by auto qed qed auto qed qed lemma croot_via_root_selection: assumes roots: "set ys = { y. y^n = x}" and n: "n \ 0" shows "croot n x = arg_min_list (\ y. (- Re y, - Im y)) ys" (is "_ = arg_min_list ?f ys") proof (rule croot_unique[OF n]) let ?y = "arg_min_list ?f ys" have rt: "croot n x ^ n = x" using n by (rule croot_power) hence "croot n x \ set ys" unfolding roots by auto hence ys: "ys \ []" by auto from arg_min_list_in[OF this] have "?y \ set ys" by auto from this[unfolded roots] show "?y^n = x" by auto fix z assume "z^n = x" hence z: "z \ set ys" unfolding roots by auto from f_arg_min_list_f[OF ys, of ?f] z have "?f ?y \ ?f z" by simp thus "Re z < Re ?y \ Re z = Re ?y \ Im z \ Im ?y" by auto qed lemma croot_impl[code]: "croot n x = (if n = 0 then 0 else arg_min_list (\ y. (- Re y, - Im y)) (all_croots n x))" proof (cases "n = 0") case n0: False hence id: "(if n = 0 then y else z) = z" for y z u :: complex by auto show ?thesis unfolding id Let_def by (rule croot_via_root_selection[OF _ n0], rule all_croots[OF n0]) qed auto end \ No newline at end of file diff --git a/thys/Cubic_Quartic_Equations/Min_Int_Poly_Complex_Impl.thy b/thys/Cubic_Quartic_Equations/Min_Int_Poly_Complex_Impl.thy deleted file mode 100644 --- a/thys/Cubic_Quartic_Equations/Min_Int_Poly_Complex_Impl.thy +++ /dev/null @@ -1,135 +0,0 @@ -section \Implementation of the minimal polynomial of a complex algebraic number\ - -theory Min_Int_Poly_Complex_Impl - imports - Algebraic_Numbers.Complex_Algebraic_Numbers -begin - -definition complex_poly :: "int poly \ int poly \ int poly list" where - "complex_poly re im = (let i = [:1,0,1:] - in factors_of_int_poly (poly_add re (poly_mult im i)))" - -lemma complex_poly: assumes re: "re represents (Re x)" - and im: "im represents (Im x)" - shows "\ f \ set (complex_poly re im). f represents x" "\ f. f \ set (complex_poly re im) \ poly_cond f" -proof - - let ?p = "poly_add re (poly_mult im [:1, 0, 1:])" - from re have re: "re represents complex_of_real (Re x)" by simp - from im have im: "im represents complex_of_real (Im x)" by simp - have "[:1,0,1:] represents \" by auto - from represents_add[OF re represents_mult[OF im this]] - have "?p represents of_real (Re x) + complex_of_real (Im x) * \" by simp - also have "of_real (Re x) + complex_of_real (Im x) * \ = x" - by (metis complex_eq mult.commute) - finally have p: "?p represents x" by auto - have "factors_of_int_poly ?p = complex_poly re im" - unfolding complex_poly_def Let_def by simp - from factors_of_int_poly(1)[OF this] factors_of_int_poly(2)[OF this, of x] p - show "\ f \ set (complex_poly re im). f represents x" "\ f. f \ set (complex_poly re im) \ poly_cond f" - unfolding represents_def by auto -qed - - -definition algebraic_real :: "real \ bool" where - [simp]: "algebraic_real = algebraic" - -lemma algebraic_real_iff[code_unfold]: "algebraic = algebraic_real" by simp - -lemma algebraic_real_code[code]: "algebraic_real (real_of x) = True" -proof (cases "info_real_alg x") - case (Inl r) - show ?thesis using info_real_alg(2)[OF Inl] by (auto simp: algebraic_of_rat) -next - case (Inr pair) - then obtain p n where Inr: "info_real_alg x = Inr (p,n)" by (cases pair, auto) - from info_real_alg(1)[OF Inr] have "p represents (real_of x)" by auto - thus ?thesis by (auto simp: algebraic_altdef_ipoly) -qed - -lemma algebraic_complex_iff[code_unfold]: "algebraic x \ (let y = x in algebraic (Re y) \ algebraic (Im y))" - unfolding Let_def -proof - assume "algebraic x" - from this[unfolded algebraic_altdef_ipoly] obtain p where "ipoly p x = 0" "p \ 0" by auto - from represents_root_poly[OF this] show "algebraic (Re x) \ algebraic (Im x)" - unfolding represents_def algebraic_altdef_ipoly by auto -next - assume "algebraic (Re x) \ algebraic (Im x)" - from this[unfolded algebraic_altdef_ipoly] obtain re im where - "re represents (Re x)" "im represents (Im x)" by blast - from complex_poly[OF this] show "algebraic x" - unfolding represents_def algebraic_altdef_ipoly by auto -qed - -lemma algebraic_0[simp]: "algebraic 0" - unfolding algebraic_altdef_ipoly - by (intro exI[of _ "[:0,1:]"], auto) - -lemma min_int_poly_complex_of_real[simp]: "min_int_poly (complex_of_real x) = min_int_poly x" -proof (cases "algebraic x") - case False - hence "\ algebraic (complex_of_real x)" unfolding algebraic_complex_iff by auto - with False show ?thesis unfolding min_int_poly_def by auto -next - case True - from min_int_poly_represents[OF True] - have "min_int_poly x represents x" by auto - thus ?thesis - by (intro min_int_poly_unique, auto simp: lead_coeff_min_int_poly_pos) -qed - - - -text \TODO: the implemention might be tuned, since the search process should be faster when - using interval arithmetic to figure out the correct factor. - (One might also implement the search via checking @{term "ipoly f x = 0"}, but because of complex-algebraic-number - arithmetic, I think that search would be slower than the current one via "@{term "x \ set (complex_roots_of_int_poly f)"}\ -definition min_int_poly_complex :: "complex \ int poly" where - "min_int_poly_complex x = (if algebraic x then if Im x = 0 then min_int_poly_real (Re x) - else the (find (\ f. x \ set (complex_roots_of_int_poly f)) (complex_poly (min_int_poly (Re x)) (min_int_poly (Im x)))) - else [:0,1:])" - -lemma min_int_poly_complex[code_unfold]: "min_int_poly = min_int_poly_complex" -proof (standard) - fix x - define fs where "fs = complex_poly (min_int_poly (Re x)) (min_int_poly (Im x))" - let ?f = "min_int_poly_complex x" - show "min_int_poly x = ?f" - proof (cases "algebraic x") - case False - thus ?thesis unfolding min_int_poly_def min_int_poly_complex_def by auto - next - case True - show ?thesis - proof (cases "Im x = 0") - case *: True - have id: "?f = min_int_poly_real (Re x)" unfolding min_int_poly_complex_def * using True by auto - show ?thesis unfolding id min_int_poly_real_code_unfold[symmetric] min_int_poly_complex_of_real[symmetric] - using * by (intro arg_cong[of _ _ min_int_poly] complex_eqI, auto) - next - case False - from True[unfolded algebraic_complex_iff] have "algebraic (Re x)" "algebraic (Im x)" by auto - from complex_poly[OF min_int_poly_represents[OF this(1)] min_int_poly_represents[OF this(2)]] - have fs: "\ f \ set fs. ipoly f x = 0" "\ f. f \ set fs \ poly_cond f" unfolding fs_def by auto - let ?fs = "find (\ f. ipoly f x = 0) fs" - let ?fs' = "find (\ f. x \ set (complex_roots_of_int_poly f)) fs" - have "?f = the ?fs'" unfolding min_int_poly_complex_def fs_def - using True False by auto - also have "?fs' = ?fs" - by (rule find_cong[OF refl], subst complex_roots_of_int_poly, insert fs, auto) - finally have id: "?f = the ?fs" . - from fs(1) have "?fs \ None" unfolding find_None_iff by auto - then obtain f where Some: "?fs = Some f" by auto - from find_Some_D[OF this] fs(2)[of f] - show ?thesis unfolding id Some - by (intro min_int_poly_unique, auto) - qed - qed -qed - -(* outcommented tests, since time-consuming: - -value (code) "min_int_poly (sqrt 2 + \)" - -*) -end diff --git a/thys/Cubic_Quartic_Equations/ROOT b/thys/Cubic_Quartic_Equations/ROOT --- a/thys/Cubic_Quartic_Equations/ROOT +++ b/thys/Cubic_Quartic_Equations/ROOT @@ -1,11 +1,12 @@ chapter AFP session Cubic_Quartic_Equations (AFP) = Algebraic_Numbers + options [timeout = 600] sessions Complex_Geometry + Factor_Algebraic_Polynomial theories Quartic_Polynomials document_files "root.tex" "root.bib" diff --git a/thys/Factor_Algebraic_Polynomial/Factor_Complex_Poly.thy b/thys/Factor_Algebraic_Polynomial/Factor_Complex_Poly.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Factor_Complex_Poly.thy @@ -0,0 +1,141 @@ +section \Factorization of Polynomials with Algebraic Coefficients\ + +subsection \Complex Algebraic Coefficients\ + +theory Factor_Complex_Poly + imports + Roots_of_Real_Complex_Poly +begin +hide_const (open) MPoly_Type.smult MPoly_Type.degree MPoly_Type.coeff MPoly_Type.coeffs + +definition factor_complex_main :: "complex poly \ complex \ (complex \ nat) list" where + "factor_complex_main p \ let (c,pis) = yun_rf (yun_polys p) in + (c, concat (map (\ (p,i). map (\ r. (r,i)) (roots_of_complex_rf_poly p)) pis))" + +lemma roots_of_complex_poly_via_factor_complex_main: + "map fst (snd (factor_complex_main p)) = roots_of_complex_poly p" +proof (cases "p = 0") + case True + have [simp]: "yun_rf (yun_polys 0) = (0,[])" + by (transfer, simp) + show ?thesis + unfolding factor_complex_main_def Let_def roots_of_complex_poly_def True + by simp +next + case False + hence p: "(p = 0) = False" by simp + obtain c rts where yun: "yun_rf (yun_polys p) = (c,rts)" by force + show ?thesis + unfolding factor_complex_main_def Let_def roots_of_complex_poly_def p if_False + roots_of_complex_rf_polys_def polys_rf_def o_def yun split snd_conv map_map + by (induct rts, auto simp: o_def) +qed + +lemma distinct_factor_complex_main: + "distinct (map fst (snd (factor_complex_main p)))" + unfolding roots_of_complex_poly_via_factor_complex_main + by (rule distinct_roots_of_complex_poly) + +lemma factor_complex_main: assumes rt: "factor_complex_main p = (c,xis)" + shows "p = smult c (\(x, i)\xis. [:- x, 1:] ^ Suc i)" +proof - + obtain d pis where yun: "yun_factorization gcd p = (d,pis)" by force + obtain d' pis' where yun_rf: "yun_rf (yun_polys p) = (d',pis')" by force + let ?p = poly_rf + let ?map = "map (\ (p,i). (?p p, i))" + from yun yun_rf have d': "d' = d" and pis: "pis = ?map pis'" + by (atomize(full), transfer, auto) + from rt[unfolded factor_complex_main_def yun_rf split Let_def d'] + have xis: "xis = concat (map (\(p, i). map (\r. (r, i)) (roots_of_complex_rf_poly p)) pis')" + and d: "d = c" + by (auto split: if_splits) + note yun = yun_factorization[OF yun[unfolded d]] + note yun = square_free_factorizationD[OF yun(1)] yun(2)[unfolded snd_conv] + let ?exp = "\ pis. \(x, i)\concat + (map (\(p, i). map (\r. (r, i)) (roots_of_complex_rf_poly p)) pis). [:- x, 1:] ^ Suc i" + from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . + also have "(\(a, i)\set pis. a ^ Suc i) = (\(a, i)\pis. a ^ Suc i)" + by (rule prod.distinct_set_conv_list[OF yun(5)]) + also have "\ = ?exp pis'" using yun(2,6) unfolding pis + proof (induct pis') + case (Cons pi pis) + obtain p i where pi: "pi = (p,i)" by force + let ?rts = "roots_of_complex_rf_poly p" + note Cons = Cons[unfolded pi] + have IH: "(\(a, i)\?map pis. a ^ Suc i) = (?exp pis)" + by (rule Cons(1)[OF Cons(2-3)], auto) + from Cons(2-3)[of "?p p" i] have p: "square_free (?p p)" "degree (?p p) \ 0" "?p p \ 0" "monic (?p p)" by auto + have "(\(a, i)\?map (pi # pis). a ^ Suc i) = ?p p ^ Suc i * (\(a, i)\?map pis. a ^ Suc i)" + unfolding pi by simp + also have "(\(a, i)\?map pis. a ^ Suc i) = ?exp pis" by (rule IH) + finally have id: "(\(a, i)\?map (pi # pis). a ^ Suc i) = ?p p ^ Suc i * ?exp pis" by simp + have "?exp (pi # pis) = ?exp [(p,i)] * ?exp pis" unfolding pi by simp + also have "?exp [(p,i)] = (\(x, i)\ (map (\r. (r, i)) ?rts). [:- x, 1:] ^ Suc i)" + by simp + also have "\ = (\ x \ ?rts. [:- x, 1:])^Suc i" + unfolding prod_list_power by (rule arg_cong[of _ _ prod_list], auto) + also have "(\ x \ ?rts. [:- x, 1:]) = ?p p" + proof - + from fundamental_theorem_algebra_factorized[of "?p p", unfolded \monic (?p p)\] + obtain as where as: "?p p = (\a\as. [:- a, 1:])" by (metis smult_1_left) + also have "\ = (\a\set as. [:- a, 1:])" + proof (rule sym, rule prod.distinct_set_conv_list, rule ccontr) + assume "\ distinct as" + from not_distinct_decomp[OF this] obtain as1 as2 as3 a where + a: "as = as1 @ [a] @ as2 @ [a] @ as3" by blast + define q where "q = (\a\as1 @ as2 @ as3. [:- a, 1:])" + have "?p p = (\a\as. [:- a, 1:])" by fact + also have "\ = (\a\([a] @ [a]). [:- a, 1:]) * q" + unfolding q_def a map_append prod_list.append by (simp only: ac_simps) + also have "\ = [:-a,1:] * [:-a,1:] * q" by simp + finally have "?p p = ([:-a,1:] * [:-a,1:]) * q" by simp + hence "[:-a,1:] * [:-a,1:] dvd ?p p" unfolding dvd_def .. + with \square_free (?p p)\[unfolded square_free_def, THEN conjunct2, rule_format, of "[:-a,1:]"] + show False by auto + qed + also have "set as = {x. poly (?p p) x = 0}" unfolding as poly_prod_list + by (simp add: o_def, induct as, auto) + also have "\ = set ?rts" by (simp add: roots_of_complex_rf_poly(1)) + also have "(\a\set ?rts. [:- a, 1:]) = (\a\?rts. [:- a, 1:])" + by (rule prod.distinct_set_conv_list[OF roots_of_complex_rf_poly(2)]) + finally show ?thesis by simp + qed + finally have id2: "?exp (pi # pis) = ?p p ^ Suc i * ?exp pis" by simp + show ?case unfolding id id2 .. + qed simp + also have "?exp pis' = (\(x, i)\xis. [:- x, 1:] ^ Suc i)" unfolding xis .. + finally show ?thesis unfolding p xis by simp +qed + +definition factor_complex_poly :: "complex poly \ complex \ (complex poly \ nat) list" where + "factor_complex_poly p = (case factor_complex_main p of + (c,ris) \ (c, map (\ (r,i). ([:-r,1:],Suc i)) ris))" + +lemma distinct_factor_complex_poly: + "distinct (map fst (snd (factor_complex_poly p)))" +proof - + obtain c ris where main: "factor_complex_main p = (c,ris)" by force + show ?thesis unfolding factor_complex_poly_def main split + using distinct_factor_complex_main[of p, unfolded main] + unfolding snd_conv o_def + unfolding distinct_map by (force simp: inj_on_def) +qed + + +theorem factor_complex_poly: assumes fp: "factor_complex_poly p = (c,qis)" + shows + "p = smult c (\(q, i)\qis. q ^ i)" + "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" +proof - + from fp[unfolded factor_complex_poly_def] + obtain pis where fp: "factor_complex_main p = (c,pis)" + and qis: "qis = map (\(r, i). ([:- r, 1:], Suc i)) pis" + by (cases "factor_complex_main p", auto) + from factor_complex_main[OF fp] have p: "p = smult c (\(x, i)\pis. [:- x, 1:] ^ Suc i)" . + show "p = smult c (\(q, i)\qis. q ^ i)" unfolding p qis + by (rule arg_cong[of _ _ "\ p. smult c (prod_list p)"], auto) + show "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" + using linear_irreducible_field[of q] unfolding qis by auto +qed + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Factor_Real_Poly.thy b/thys/Factor_Algebraic_Polynomial/Factor_Real_Poly.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Factor_Real_Poly.thy @@ -0,0 +1,288 @@ +(* Factorization of polynomials with real algebraic coefficients *) + +subsection \Real Algebraic Coefficients\ + +text \We basically perform a factorization via complex algebraic numbers, + take all real roots, and + then merge each pair of conjugate roots into a quadratic factor.\ + +theory Factor_Real_Poly + imports + Factor_Complex_Poly +begin + +hide_const (open) Coset.order + +fun delete_cnj :: "complex \ nat \ (complex \ nat) list \ (complex \ nat) list" where + "delete_cnj x i ((y,j) # yjs) = (if x = y then if Suc j = i then yjs else if Suc j > i then + ((y,j - i) # yjs) else delete_cnj x (i - Suc j) yjs else (y,j) # delete_cnj x i yjs)" +| "delete_cnj _ _ [] = []" + +lemma delete_cnj_length[termination_simp]: "length (delete_cnj x i yjs) \ length yjs" + by (induct x i yjs rule: delete_cnj.induct, auto) + +fun complex_roots_to_real_factorization :: "(complex \ nat) list \ (real poly \ nat)list" where + "complex_roots_to_real_factorization [] = []" +| "complex_roots_to_real_factorization ((x,i) # xs) = (if x \ \ then + ([:-(Re x),1:],Suc i) # complex_roots_to_real_factorization xs else + let xx = cnj x; ys = delete_cnj xx (Suc i) xs; p = map_poly Re ([:-x,1:] * [:-xx,1:]) + in (p,Suc i) # complex_roots_to_real_factorization ys)" + +definition factor_real_poly :: "real poly \ real \ (real poly \ nat) list" where + "factor_real_poly p \ case factor_complex_main (map_poly of_real p) of + (c,ris) \ (Re c, complex_roots_to_real_factorization ris) " + + +lemma monic_imp_nonzero: "monic x \ x \ 0" for x :: "'a :: semiring_1 poly" by auto + +lemma delete_cnj: assumes + "order x (\(x, i)\xis. [:- x, 1:] ^ Suc i) \ si" "si \ 0" + shows "(\(x, i)\xis. [:- x, 1:] ^ Suc i) = + [:- x, 1:] ^ si * (\(x, i)\delete_cnj x si xis. [:- x, 1:] ^ Suc i)" +using assms +proof (induct x si xis rule: delete_cnj.induct) + case (2 x si) + hence "order x 1 \ 1" by auto + hence "[:-x,1:]^1 dvd 1" unfolding order_divides by simp + from power_le_dvd[OF this, of 1] \si \ 0\ have "[:- x, 1:] dvd 1" by simp + from divides_degree[OF this] + show ?case by auto +next + case (1 x i y j yjs) + note IH = 1(1-2) + let ?yj = "[:-y,1:]^Suc j" + let ?yjs = "(\(x,i)\yjs. [:- x, 1:] ^ Suc i)" + let ?x = "[: - x, 1 :]" + let ?xi = "?x ^ i" + have "monic (\(x,i)\(y, j) # yjs. [:- x, 1:] ^ Suc i)" + by (rule monic_prod_list_pow) then have "monic (?yj * ?yjs)" by simp + from monic_imp_nonzero[OF this] have yy0: "?yj * ?yjs \ 0" by auto + have id: "(\(x,i)\(y, j) # yjs. [:- x, 1:] ^ Suc i) = ?yj * ?yjs" by simp + from 1(3-) have ord: "i \ order x (?yj * ?yjs)" and i: "i \ 0" unfolding id by auto + from ord[unfolded order_mult[OF yy0]] have ord: "i \ order x ?yj + order x ?yjs" . + from this[unfolded order_linear_power'] + have ord: "i \ (if y = x then Suc j else 0) + order x ?yjs" by simp + show ?case + proof (cases "x = y") + case False + from ord False have "i \ order x ?yjs" by simp + note IH = IH(2)[OF False this i] + from False have del: "delete_cnj x i ((y, j) # yjs) = (y,j) # delete_cnj x i yjs" by simp + show ?thesis unfolding del id IH + by (simp add: ac_simps) + next + case True note xy = this + note IH = IH(1)[OF True] + show ?thesis + proof (cases "Suc j \ i") + case False + from ord have ord: "i - Suc j \ order x ?yjs" unfolding xy by simp + have "?xi = ?x ^ (Suc j + (i - Suc j))" using False by simp + also have "\ = ?x ^ Suc j * ?x ^ (i - Suc j)" + unfolding power_add by simp + finally have xi: "?xi = ?x ^ Suc j * ?x ^ (i - Suc j)" . + from False have "Suc j \ i" "\ i < Suc j" "i - Suc j \ 0" by auto + note IH = IH[OF this(1,2) ord this(3)] + from xy False have del: "delete_cnj x i ((y, j) # yjs) = delete_cnj x (i - Suc j) yjs" by auto + show ?thesis unfolding del id unfolding IH xi unfolding xy by simp + next + case True + hence "Suc j = i \ i < Suc j" by auto + thus ?thesis + proof + assume i: "Suc j = i" + from xy i have del: "delete_cnj x i ((y, j) # yjs) = yjs" by simp + show ?thesis unfolding id del unfolding xy i by simp + next + assume ij: "i < Suc j" + with xy i have del: "delete_cnj x i ((y, j) # yjs) = (y, j - i) # yjs" by simp + from ij have idd: "Suc j = i + Suc (j - i)" by simp + show ?thesis unfolding id del unfolding xy idd power_add by simp + qed + qed + qed +qed + + +theorem factor_real_poly: assumes fp: "factor_real_poly p = (c,qis)" + shows "p = smult c (\(q, i)\qis. q ^ i)" + "(q,j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1,2}" +proof - + interpret map_poly_inj_idom_hom of_real.. + have "(p = smult c (\(q, i)\qis. q ^ i)) \ ((q,j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1,2})" + proof (cases "p = 0") + case True + have yun: "yun_rf (yun_polys (0 :: complex poly)) = (0,[])" + by (transfer, auto simp: yun_factorization_def) + have "factor_real_poly p = (0,[])" unfolding True + by (simp add: factor_real_poly_def factor_complex_main_def yun) + with fp have id: "c = 0" "qis = []" by auto + thus ?thesis unfolding True by simp + next + case False note p0 = this + let ?c = complex_of_real + let ?rp = "map_poly Re" + let ?cp = "map_poly ?c" + let ?p = "?cp p" + from fp[unfolded factor_real_poly_def] + obtain d xis where fp: "factor_complex_main ?p = (d,xis)" + and c: "c = Re d" and qis: "qis = complex_roots_to_real_factorization xis" + by (cases "factor_complex_main ?p", auto) + from factor_complex_main[OF fp] have p: "?p = smult d (\(x, i)\xis. [:- x, 1:] ^ Suc i)" + (is "_ = smult d ?q") . + from arg_cong[OF this, of "\ p. coeff p (degree p)"] + have "coeff ?p (degree ?p) = coeff (smult d ?q) (degree (smult d ?q))" . + also have "coeff ?p (degree ?p) = ?c (coeff p (degree p))" by simp + also have "coeff (smult d ?q) (degree (smult d ?q)) = d * coeff ?q (degree ?q)" + by simp + also have "monic ?q" by (rule monic_prod_list_pow) + finally have d: "d = ?c (coeff p (degree p))" by auto + from arg_cong[OF this, of Re, folded c] have c: "c = coeff p (degree p)" by auto + have "set (coeffs ?p) \ \" by auto + with p have q': "set (coeffs (smult d ?q)) \ \" by auto + from d p0 have d0: "d \ 0" by auto + have "smult d ?q = [:d:] * ?q" by auto + from real_poly_factor[OF q'[unfolded this]] d0 d + have q: "set (coeffs ?q) \ \" by auto + have "p = ?rp ?p" + by (rule sym, subst map_poly_map_poly, force, rule map_poly_idI, auto) + also have "\ = ?rp (smult d ?q)" unfolding p .. + also have "?q = ?cp (?rp ?q)" + by (rule sym, rule map_poly_of_real_Re, insert q, auto) + also have "d = ?c c" unfolding d c .. + also have "smult (?c c) (?cp (?rp ?q)) = ?cp (smult c (?rp ?q))" by (simp add: hom_distribs) + also have "?rp \ = smult c (?rp ?q)" + by (subst map_poly_map_poly, force, rule map_poly_idI, auto) + finally have p: "p = smult c (?rp ?q)" . + let ?fact = complex_roots_to_real_factorization + have "?rp ?q = (\(q, i)\qis. q ^ i) \ + ((q, j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1, 2})" + using q unfolding qis + proof (induct xis rule: complex_roots_to_real_factorization.induct) + case 1 + show ?case by simp + next + case (2 x i xis) + note IH = 2(1-2) + note prems = 2(3) + let ?xi = "[:- x, 1:] ^ Suc i" + let ?xis = "(\(x, i)\xis. [:- x, 1:] ^ Suc i)" + have id: "(\(x, i)\((x,i) # xis). [:- x, 1:] ^ Suc i) = ?xi * ?xis" + by simp + show ?case + proof (cases "x \ \") + case True + have xi: "set (coeffs ?xi) \ \" + by (rule real_poly_power, insert True, auto) + have xis: "set (coeffs ?xis) \ \" + by (rule real_poly_factor[OF prems[unfolded id] xi], rule linear_power_nonzero) + note IH = IH(1)[OF True xis] + have "?rp (?xi * ?xis) = ?rp ?xi * ?rp ?xis" + by (rule map_poly_Re_mult[OF xi xis]) + also have "?rp ?xi = (?rp [: -x,1 :])^Suc i" + by (rule map_poly_Re_power, insert True, auto) + also have "?rp [: -x,1 :] = [:-(Re x),1:]" by auto + also have "?rp ?xis = (\ (a,b) \ ?fact xis. a ^ b)" + using IH by auto + also have "[:- Re x, 1:] ^ Suc i * (\ (a,b) \ ?fact xis. a ^ b) = + (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" using True by simp + finally have idd: "?rp (?xi * ?xis) = (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" . + show ?thesis unfolding id idd + proof (intro conjI, force, intro impI) + assume "(q, j) \ set (?fact ((x, i) # xis))" + hence "(q,j) \ set (?fact xis) \ (q = [:- Re x, 1:] \ j = Suc i)" + using True by auto + thus "irreducible q \ j \ 0 \ monic q \ degree q \ {1, 2}" + proof + assume "(q,j) \ set (?fact xis)" + with IH show ?thesis by auto + next + assume "q = [:- Re x, 1:] \ j = Suc i" + with linear_irreducible_field[of "[:- Re x, 1:]"] show ?thesis by auto + qed + qed + next + case False + define xi where "xi = [:Re x * Re x + Im x * Im x, - (2 * Re x), 1:]" + obtain xx where xx: "xx = cnj x" by auto + have xi: "xi = ?rp ([:-x,1:] * [:-xx,1:])" unfolding xx xi_def by auto + have cpxi: "?cp xi = [:-x,1:] * [:-xx,1:]" unfolding xi_def + by (cases x, auto simp: xx legacy_Complex_simps) + obtain yis where yis: "yis = delete_cnj xx (Suc i) xis" by auto + from False have fact: "?fact ((x,i) # xis) = ((xi,Suc i) # ?fact yis)" + unfolding xi_def xx yis by simp + note IH = IH(2)[OF False xx yis xi] + have "irreducible xi" + apply (fold irreducible_connect_field) + proof (rule irreducible\<^sub>dI) + show "degree xi > 0" unfolding xi by auto + fix q p :: "real poly" + assume "degree q > 0" "degree q < degree xi" and qp: "xi = q * p" + hence dq: "degree q = 1" unfolding xi by auto + have dxi: "degree xi = 2" "xi \ 0" unfolding xi by auto + with qp have "q \ 0" "p \ 0" by auto + hence "degree xi = degree q + degree p" unfolding qp + by (rule degree_mult_eq) + with dq have dp: "degree p = 1" unfolding dxi by simp + { + fix c :: complex + assume rt: "poly (?cp xi) c = 0" + hence "poly (?cp q * ?cp p) c = 0" by (simp add: qp hom_distribs) + hence "(poly (?cp q) c = 0 \ poly (?cp p) c = 0)" by auto + hence "c = roots1 (?cp q) \ c = roots1 (?cp p)" + using roots1[of "?cp q"] roots1[of "?cp p"] dp dq by auto + hence "c \ \" unfolding roots1_def by auto + hence "c \ x" using False by auto + } + hence "poly (?cp xi) x \ 0" by auto + thus False unfolding cpxi by simp + qed + hence xi': "irreducible xi" "monic xi" "degree xi = 2" + unfolding xi by auto + let ?xxi = "[:- xx, 1:] ^ Suc i" + let ?yis = "(\(x, i)\yis. [:- x, 1:] ^ Suc i)" + let ?yi = "(?cp xi)^Suc i" + have yi: "set (coeffs ?yi) \ \" + by (rule real_poly_power, auto simp: xi) + have mon: "monic (\(x, i)\(x, i) # xis. [:- x, 1:] ^ Suc i)" + by (rule monic_prod_list_pow) + from monic_imp_nonzero[OF this] have xixis: "?xi * ?xis \ 0" unfolding id by auto + from False have xxx: "xx \ x" unfolding xx by (cases x, auto simp: legacy_Complex_simps Reals_def) + from prems[unfolded id] have prems: "set (coeffs (?xi * ?xis)) \ \" . + from id have "[:- x, 1:] ^ Suc i dvd ?xi * ?xis" by auto + from xixis this[unfolded order_divides] + have "order x (?xi * ?xis) \ Suc i" by auto + with complex_conjugate_order[OF prems xixis, of x, folded xx] + have "order xx (?xi * ?xis) \ Suc i" by auto + hence "order xx ?xi + order xx ?xis \ Suc i" unfolding order_mult[OF xixis] . + also have "order xx ?xi = 0" unfolding order_linear_power' using xxx by simp + finally have "order xx ?xis \ Suc i" by simp + hence yis: "?xis = ?xxi * ?yis" unfolding yis + by (rule delete_cnj, simp) + hence "?xi * ?xis = (?xi * ?xxi) * ?yis" by (simp only: ac_simps) + also have "?xi * ?xxi = ([:- x, 1:] * [:- xx, 1:])^Suc i" + by (metis power_mult_distrib) + also have "[:- x, 1:] * [:- xx, 1:] = ?cp xi" unfolding cpxi .. + finally have idd: "?xi * ?xis = (?cp xi)^Suc i * ?yis" by simp + from prems[unfolded idd] have R: "set (coeffs ((?cp xi)^Suc i * ?yis)) \ \" . + have yis: "set (coeffs ?yis) \ \" + by (rule real_poly_factor[OF R yi], auto, auto simp: xi_def) + note IH = IH[OF yis] + have "?rp (?xi * ?xis) = ?rp ?yi * ?rp ?yis" unfolding idd + by (rule map_poly_Re_mult[OF yi yis]) + also have "?rp ?yi = xi^Suc i" by (fold hom_distribs, rule map_poly_Re_of_real) + also have "?rp ?yis = (\ (a,b) \ ?fact yis. a ^ b)" + using IH by auto + also have "xi ^ Suc i * (\ (a,b) \ ?fact yis. a ^ b) = + (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" unfolding fact by simp + finally have idd: "?rp (?xi * ?xis) = (\ (a,b) \ ?fact ((x,i) # xis). a ^ b)" . + show ?thesis unfolding id idd fact using IH xi' by auto + qed + qed + thus ?thesis unfolding p by simp + qed + thus "p = smult c (\(q, i)\qis. q ^ i)" + "(q,j) \ set qis \ irreducible q \ j \ 0 \ monic q \ degree q \ {1,2}" by blast+ +qed + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Is_Int_To_Int.thy b/thys/Factor_Algebraic_Polynomial/Is_Int_To_Int.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Is_Int_To_Int.thy @@ -0,0 +1,44 @@ +section \Testing for Integrality and Conversion to Integers\ + +theory Is_Int_To_Int + imports + Polynomial_Interpolation.Is_Rat_To_Rat +begin + +lemma inv_of_rat: "inv of_rat (of_rat x) = x" + by (meson injI inv_f_eq of_rat_eq_iff) + +lemma of_rat_Ints_iff: "((of_rat x :: 'a :: field_char_0) \ \) = (x \ \)" + by (metis Ints_cases Ints_of_int inv_of_rat of_rat_of_int_eq) + +lemma is_int_code[code_unfold]: + shows "(x \ \) = (is_rat x \ is_int_rat (to_rat x))" +proof - + have "x \ \ \ x \ \ \ x \ \" + by (metis Ints_cases Rats_of_int) + also have "\ = (is_rat x \ is_int_rat (to_rat x))" + proof (simp, intro conj_cong[OF refl]) + assume "x \ \" + then obtain y where x: "x = of_rat y" unfolding Rats_def by auto + show "(x \ \) = (to_rat x \ \)" unfolding x + by (simp add: of_rat_Ints_iff) + qed + finally show ?thesis . +qed + +definition to_int :: "'a :: is_rat \ int" where + "to_int x = int_of_rat (to_rat x)" + +lemma of_int_to_int: "x \ \ \ of_int (to_int x) = x" + by (metis Ints_cases int_of_rat(1) of_rat_of_int_eq to_int_def to_rat_of_rat) + +lemma to_int_of_int: "to_int (of_int x) = x" + by (metis int_of_rat(1) of_rat_of_int_eq to_int_def to_rat_of_rat) + +lemma to_rat_complex_of_real[simp]: "to_rat (complex_of_real x) = to_rat x" + by (metis Re_complex_of_real complex_of_real_of_rat of_rat_to_rat to_rat to_rat_of_rat) + +lemma to_int_complex_of_real[simp]: "to_int (complex_of_real x) = to_int x" + by (simp add: to_int_def) + +end diff --git a/thys/Factor_Algebraic_Polynomial/MPoly_Container.thy b/thys/Factor_Algebraic_Polynomial/MPoly_Container.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/MPoly_Container.thy @@ -0,0 +1,16 @@ +subsection \Class Instances for Multivariate Polynomials and Containers\ + + +theory MPoly_Container + imports + "Polynomials.MPoly_Type_Class" + "Containers.Set_Impl" +begin + +text \Basic setup for using multivariate polynomials in combination with container framework.\ + +derive (eq) ceq poly_mapping +derive (dlist) set_impl poly_mapping (* difference list *) +derive (no) ccompare poly_mapping (* no order on poly-mapping *) + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/MPoly_Divide.thy b/thys/Factor_Algebraic_Polynomial/MPoly_Divide.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/MPoly_Divide.thy @@ -0,0 +1,295 @@ +subsection \Exact Division of Multivariate Polynomials\ + +theory MPoly_Divide + imports + Hermite_Lindemann.More_Multivariate_Polynomial_HLW + Polynomials.MPoly_Type_Class + Poly_Connection +begin + +lemma poly_lead_coeff_dvd_lead_coeff: + assumes "p dvd (q :: 'a :: idom poly)" + shows "Polynomial.lead_coeff p dvd Polynomial.lead_coeff q" + using assms by (elim dvdE) (auto simp: Polynomial.lead_coeff_mult) + + +text \ + Since there is no particularly sensible algorithm for division with a remainder on + multivariate polynomials, we define the following division operator that performs an + exact division if possible and returns 0 otherwise. +\ + +no_notation MPoly_Type.div (infixl "div" 70) + +instantiation mpoly :: ("comm_semiring_1") divide +begin + +definition divide_mpoly :: "'a mpoly \ 'a mpoly \ 'a mpoly" where + "divide_mpoly x y = (if y \ 0 \ y dvd x then THE z. x = y * z else 0)" + +instance .. + +end + +instance mpoly :: (idom) idom_divide + by standard (auto simp: divide_mpoly_def) + + + +lemma (in transfer_mpoly_to_mpoly_poly) transfer_div [transfer_rule]: + assumes [transfer_rule]: "R p' p" "R q' q" + assumes "q dvd p" + shows "R (p' div q') (p div q)" + using assms + by (smt (z3) div_by_0 dvd_imp_mult_div_cancel_left mpoly_to_mpoly_poly_0 mpoly_to_mpoly_poly_eq_iff + mpoly_to_mpoly_poly_mult nonzero_mult_div_cancel_left transfer_mpoly_to_mpoly_poly.R_def) + + +instantiation mpoly :: ("{normalization_semidom, idom}") normalization_semidom +begin + +definition unit_factor_mpoly :: "'a mpoly \ 'a mpoly" where + "unit_factor_mpoly p = Const (unit_factor (lead_coeff p))" + +definition normalize_mpoly :: "'a mpoly \ 'a mpoly" where + "normalize_mpoly p = Rings.divide p (unit_factor p)" + +lemma unit_factor_mpoly_Const [simp]: + "unit_factor (Const c) = Const (unit_factor c)" + unfolding unit_factor_mpoly_def by simp + +lemma normalize_mpoly_Const [simp]: + "normalize (Const c) = Const (normalize c)" +proof (cases "c = 0") + case False + have "normalize (Const c) = Const c div Const (unit_factor c)" + by (simp add: normalize_mpoly_def) + also have "\ = Const (unit_factor c * normalize c) div Const (unit_factor c)" + by simp + also have "\ = Const (unit_factor c) * Const (normalize c) div Const (unit_factor c)" + by (subst mpoly_Const_mult) auto + also have "\ = Const (normalize c)" + using \c \ 0\ + by (subst nonzero_mult_div_cancel_left) auto + finally show ?thesis . +qed (auto simp: normalize_mpoly_def) + +instance proof + show "unit_factor (0 :: 'a mpoly) = 0" + by (simp add: unit_factor_mpoly_def) +next + show "unit_factor x = x" if "x dvd 1" for x :: "'a mpoly" + using that by (auto elim!: mpoly_is_unitE simp: is_unit_unit_factor) +next + fix x :: "'a mpoly" + assume "x \ 0" + thus "unit_factor x dvd 1" + by (auto simp: unit_factor_mpoly_def) +next + fix x y :: "'a mpoly" + assume "x dvd 1" + hence "x \ 0" + by auto + show "unit_factor (x * y) = x * unit_factor y" + proof (cases "y = 0") + case False + have "Const (unit_factor (lead_coeff x * lead_coeff y)) = + x * Const (unit_factor (lead_coeff y))" using \x dvd 1\ + by (subst unit_factor_mult_unit_left) + (auto elim!: mpoly_is_unitE simp: mpoly_Const_mult) + thus ?thesis using \x \ 0\ False + by (simp add: unit_factor_mpoly_def lead_coeff_mult) + qed (auto simp: unit_factor_mpoly_def) +next + fix p :: "'a mpoly" + let ?c = "Const (unit_factor (lead_coeff p))" + show "unit_factor p * normalize p = p" + proof (cases "p = 0") + case False + hence "?c dvd 1" + by (intro is_unit_ConstI) auto + also have "1 dvd p" + by simp + finally have "?c * (p div ?c) = p" + by (rule dvd_imp_mult_div_cancel_left) + thus ?thesis + by (auto simp: unit_factor_mpoly_def normalize_mpoly_def) + qed (auto simp: normalize_mpoly_def) +next + show "normalize (0 :: 'a mpoly) = 0" + by (simp add: normalize_mpoly_def) +qed + +end + + + +text \ + The following is an exact division operator that can fail, i.e.\ if the divisor does not + divide the dividend, it returns \<^const>\None\. +\ + +definition divide_option :: "'a :: idom_divide \ 'a \ 'a option" (infixl "div?" 70) where + "divide_option p q = (if q dvd p then Some (p div q) else None)" + +text \ + We now show that exact division on the ring $R[X_1,\ldots, X_n]$ can be reduced to + exact division on the ring $R[X_1,\ldots,X_n][X]$, i.e.\ we can go from \<^typ>\'a mpoly\ to + a \<^typ>\'a mpoly poly\ where the coefficients have one variable less than the original + multivariate polynomial. We basically simply use the isomorphism between these two rings. +\ + +lemma divide_option_mpoly: + fixes p q :: "'a :: idom_divide mpoly" + shows "p div? q = (let V = vars p \ vars q in + (if V = {} then + let a = MPoly_Type.coeff p 0; b = MPoly_Type.coeff q 0; c = a div b + in if b * c = a then Some (Const c) else None + else + let x = Max V; + p' = mpoly_to_mpoly_poly x p; q' = mpoly_to_mpoly_poly x q + in case p' div? q' of + None \ None + | Some r \ Some (poly r (Var x))))" (is "_ = ?rhs") +proof - + define x where "x = Max (vars p \ vars q)" + define p' where "p' = mpoly_to_mpoly_poly x p" + define q' where "q' = mpoly_to_mpoly_poly x q" + interpret transfer_mpoly_to_mpoly_poly x . + have [transfer_rule]: "R p' p" "R q' q" + by (auto simp: p'_def q'_def R_def) + + show ?thesis + proof (cases "vars p \ vars q = {}") + case True + define a where "a = MPoly_Type.coeff p 0" + define b where "b = MPoly_Type.coeff q 0" + have [simp]: "p = Const a" "q = Const b" + using True by (auto elim!: vars_emptyE simp: a_def b_def mpoly_coeff_Const) + show ?thesis + apply (cases "b = 0") + apply (auto simp: Let_def mpoly_coeff_Const mpoly_Const_mult divide_option_def elim!: dvdE) + by (metis dvd_triv_left) + next + case False + have "?rhs = + (case p' div? q' of None \ None + | Some r \ Some (poly r (Var x)))" + using False + unfolding Let_def + apply (simp only: ) + apply (subst if_False) + apply (simp flip: x_def p'_def q'_def cong: option.case_cong) + done + also have "\ = (if q' dvd p' then Some (poly (p' div q') (Var x)) else None)" + using False by (auto simp: divide_option_def) + also have "\ = p div? q" + unfolding divide_option_def + proof (intro if_cong refl arg_cong[where f = Some]) + show "(q' dvd p') = (q dvd p)" + by transfer_prover + next + assume [transfer_rule]: "q dvd p" + have "R (p' div q') (p div q)" + by transfer_prover + thus "poly (p' div q') (Var x) = p div q" + by (simp add: R_def poly_mpoly_to_mpoly_poly) + qed + finally show ?thesis .. + qed +qed + +text \ + Next, we show that exact division on the ring $R[X_1,\ldots,X_n][Y]$ can be reduced to + exact division on the ring $R[X_1,\ldots,X_n]$. This is essentially just polynomial division. +\ +lemma divide_option_mpoly_poly: + fixes p q :: "'a :: idom_divide mpoly poly" + shows "p div? q = + (if p = 0 then Some 0 + else if q = 0 then None + else let dp = Polynomial.degree p; dq = Polynomial.degree q + in if dp < dq then None + else case Polynomial.lead_coeff p div? Polynomial.lead_coeff q of + None \ None + | Some c \ ( + case (p - Polynomial.monom c (dp - dq) * q) div? q of + None \ None + | Some r \ Some (Polynomial.monom c (dp - dq) + r)))" + (is "_ = ?rhs") +proof (cases "p = 0"; cases "q = 0") + assume [simp]: "p \ 0" "q \ 0" + define dp where "dp = Polynomial.degree p" + define dq where "dq = Polynomial.degree q" + define cp where "cp = Polynomial.lead_coeff p" + define cq where "cq = Polynomial.lead_coeff q" + define mon where "mon = Polynomial.monom (cp div cq) (dp - dq)" + show ?thesis + proof (cases "dp < dq") + case True + hence "\q dvd p" + unfolding dp_def dq_def + by (meson \p \ 0\ divides_degree leD) + thus ?thesis + using True by (simp add: divide_option_def dp_def dq_def) + next + case deg: False + show ?thesis + proof (cases "cq dvd cp") + case False + hence "\q dvd p" + unfolding cq_def cp_def using poly_lead_coeff_dvd_lead_coeff by blast + thus ?thesis + using deg False by (simp add: dp_def dq_def Let_def divide_option_def cp_def cq_def) + next + case dvd1: True + show ?thesis + proof (cases "q dvd (p - mon * q)") + case False + hence "\q dvd p" + by (meson dvd_diff dvd_triv_right) + thus ?thesis + using deg dvd1 False + by (simp add: dp_def dq_def Let_def divide_option_def cp_def cq_def mon_def) + next + case dvd2: True + hence "q dvd p" + by (metis diff_eq_eq dvd_add dvd_triv_right) + have "?rhs = Some (mon + (p - mon * q) div q)" + using deg dvd1 dvd2 + by (simp add: dp_def dq_def Let_def divide_option_def cp_def cq_def mon_def) + also have "mon + (p - mon * q) div q = p div q" + using dvd2 by (elim dvdE) (auto simp: algebra_simps) + also have "Some \ = p div? q" + using \q dvd p\ by (simp add: divide_option_def) + finally show ?thesis .. + qed + qed + qed +qed (auto simp: divide_option_def) + +text \ + These two equations now serve as two mutually recursive code equations that allow us to + reduce exact division of multivariate polynomials to exact division of their coefficients. + Termination of these code equations is not shown explicitly, but is obvious since one variable + is eliminated in every step. +\ + +definition divide_option_mpoly :: "'a :: idom_divide mpoly \ _" + where "divide_option_mpoly = divide_option" + +definition divide_option_mpoly_poly :: "'a :: idom_divide mpoly poly \ _" + where "divide_option_mpoly_poly = divide_option" + +lemmas divide_option_mpoly_code [code] = + divide_option_mpoly [folded divide_option_mpoly_def divide_option_mpoly_poly_def] + +lemmas divide_option_mpoly_poly_code [code] = + divide_option_mpoly_poly [folded divide_option_mpoly_def divide_option_mpoly_poly_def] + +lemma divide_mpoly_code [code]: + fixes p q :: "'a :: idom_divide mpoly" + shows "p div q = (case divide_option_mpoly p q of None \ 0 | Some r \ r)" + by (auto simp: divide_option_mpoly_def divide_option_def divide_mpoly_def) + +end diff --git a/thys/Factor_Algebraic_Polynomial/MPoly_Divide_Code.thy b/thys/Factor_Algebraic_Polynomial/MPoly_Divide_Code.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/MPoly_Divide_Code.thy @@ -0,0 +1,175 @@ +subsection \Implementation of Division on Multivariate Polynomials\ + +theory MPoly_Divide_Code + imports + MPoly_Divide + Polynomials.MPoly_Type_Class_FMap + Polynomials.MPoly_Type_Univariate +begin + +text \ + We now set up code equations for some of the operations that we will need, such as division, + \<^const>\mpoly_to_poly\, and \<^const>\mpoly_to_mpoly_poly\. +\ + +lemma mapping_of_MPoly[code]: "mapping_of (MPoly p) = p" + by (simp add: MPoly_inverse) + +lift_definition filter_pm :: "('a \ bool) \ ('a \\<^sub>0 'b :: zero) \ ('a \\<^sub>0 'b)" is + "\P f x. if P x then f x else 0" + by (erule finite_subset[rotated]) auto + +lemma lookup_filter_pm: "lookup (filter_pm P f) x = (if P x then lookup f x else 0)" + by transfer auto + + +lemma filter_pm_code [code]: "filter_pm P (Pm_fmap m) = Pm_fmap (fmfilter P m)" + by (auto intro!: poly_mapping_eqI simp: fmlookup_default_def lookup_filter_pm) + +lemma remove_key_conv_filter_pm [code]: "remove_key x m = filter_pm (\y. y \ x) m" + by transfer auto + +lemma finite_poly_coeff_nonzero: "finite {n. poly.coeff p n \ 0}" + by (metis MOST_coeff_eq_0 eventually_cofinite) + +lemma poly_degree_conv_Max: + assumes "p \ 0" + shows "Polynomial.degree p = Max {n. poly.coeff p n \ 0}" + using assms +proof (intro antisym degree_le Max.boundedI) + fix n assume "n \ {n. poly.coeff p n \ 0}" + thus "n \ Polynomial.degree p" + by (simp add: le_degree) +qed (auto simp: poly_eq_iff finite_poly_coeff_nonzero) + +lemma mpoly_to_poly_code_aux: + fixes p :: "'a :: comm_monoid_add mpoly" and x :: nat + defines "I \ (\m. lookup m x) ` Set.filter (\m. \y\keys m. y = x) (keys (mapping_of p))" + shows "I = {n. poly.coeff (mpoly_to_poly x p) n \ 0}" + and "mpoly_to_poly x p = 0 \ I = {}" + and "I \ {} \ Polynomial.degree (mpoly_to_poly x p) = Max I" +proof - + have "n \ I \ poly.coeff (mpoly_to_poly x p) n \ 0" for n + proof - + have "I = (\m. lookup m x) ` (keys (mapping_of p) \ {m. \y\keys m. y = x})" + by (auto simp: I_def Set.filter_def) + also have "{m. \y\keys m. y = x} = range (\n. monomial n x)" (is "?lhs = ?rhs") + proof (intro equalityI subsetI) + fix m assume "m \ ?lhs" + hence "m = monomial (lookup m x) x" + by transfer (auto simp: fun_eq_iff when_def) + thus "m \ ?rhs" by auto + qed (auto split: if_splits) + also have "n \ (\m. lookup m x) ` (keys (mapping_of p) \ \) \ + monomial n x \ keys (mapping_of p)" by force + also have "\ \ poly.coeff (mpoly_to_poly x p) n \ 0" + by (simp add: coeff_def in_keys_iff) + finally show ?thesis . + qed + thus I: "I = {n. poly.coeff (mpoly_to_poly x p) n \ 0}" + by blast + show eq_0_iff: "mpoly_to_poly x p = 0 \ I = {}" + unfolding I by (auto simp: poly_eq_iff) + show "I \ {} \ Polynomial.degree (mpoly_to_poly x p) = Max I" + by (subst poly_degree_conv_Max) (use eq_0_iff I in auto) +qed + + +lemma mpoly_to_poly_code [code]: + "Polynomial.coeffs (mpoly_to_poly x p) = + (let I = (\m. lookup m x) ` Set.filter (\m. \y\keys m. y = x) (keys (mapping_of p)) + in if I = {} then [] else map (\n. MPoly_Type.coeff p (Poly_Mapping.single x n)) [0..m. lookup m x) ` Set.filter (\m. \y\keys m. y = x) (keys (mapping_of p))" + show ?thesis + proof (cases "I = {}") + case True + thus ?thesis using mpoly_to_poly_code_aux(2)[of x p] + by (simp add: I_def) + next + case False + have [simp]: "mpoly_to_poly x p \ 0" + using mpoly_to_poly_code_aux(2)[of x p] False by (simp add: I_def) + from False have "?rhs = map (\n. MPoly_Type.coeff p (Poly_Mapping.single x n)) [0.. = ?lhs" + proof (rule nth_equalityI) + show "length ?rhs' = length ?lhs" + using mpoly_to_poly_code_aux(3)[of x p] False + by (simp add: I_def length_coeffs_degree) + thus "?rhs' ! n = ?lhs ! n" if "n < length ?rhs'" for n using that + by (auto simp del: upt_Suc simp: nth_coeffs_coeff) + qed + finally show ?thesis .. + qed +qed + + +fun mpoly_to_mpoly_poly_impl_aux1 :: "nat \ ((nat \\<^sub>0 nat) \ 'a) list \ nat \ ((nat \\<^sub>0 nat) \ 'a) list" where + "mpoly_to_mpoly_poly_impl_aux1 i [] j = []" +| "mpoly_to_mpoly_poly_impl_aux1 i ((mon', c) # xs) j = + (if lookup mon' i = j then [(remove_key i mon', c)] else []) @ mpoly_to_mpoly_poly_impl_aux1 i xs j" + +lemma mpoly_to_mpoly_poly_impl_aux1_altdef: + "mpoly_to_mpoly_poly_impl_aux1 i xs j = + map (\(mon, c). (remove_key i mon, c)) (filter (\(mon, c). lookup mon i = j) xs)" + by (induction xs) auto + +lemma map_of_mpoly_to_mpoly_poly_impl_aux1: + "map_of (mpoly_to_mpoly_poly_impl_aux1 i xs j) = (\mon. + (if lookup mon i > 0 then None + else map_of xs (mon + Poly_Mapping.single i j)))" + apply (rule ext) + apply (induction i xs j rule: mpoly_to_mpoly_poly_impl_aux1.induct) + apply (auto simp: remove_key_lookup) + apply (meson remove_key_sum) + apply (metis add_left_cancel lookup_single_eq remove_key_sum) + apply (metis remove_key_add remove_key_single remove_key_sum single_zero) + done + +lemma lookup0_fmap_of_list_mpoly_to_mpoly_poly_impl_aux1: + "lookup0 (fmap_of_list (mpoly_to_mpoly_poly_impl_aux1 i xs j)) = (\mon. + lookup0 (fmap_of_list xs) (mon + Poly_Mapping.single i j) when lookup mon i = 0)" + by (auto simp add: fmlookup_default_def fmlookup_of_list map_of_mpoly_to_mpoly_poly_impl_aux1) + +definition mpoly_to_mpoly_poly_impl_aux2 where + "mpoly_to_mpoly_poly_impl_aux2 i p j = poly.coeff (mpoly_to_mpoly_poly i p) j" + +lemma coeff_MPoly: "MPoly_Type.coeff (MPoly f) m = lookup f m" + by (simp add: coeff_def mpoly.MPoly_inverse) + +lemma mpoly_to_mpoly_poly_impl_aux2_code [code]: + "mpoly_to_mpoly_poly_impl_aux2 i (MPoly (Pm_fmap (fmap_of_list xs))) j = + MPoly (Pm_fmap (fmap_of_list (mpoly_to_mpoly_poly_impl_aux1 i xs j)))" + unfolding mpoly_to_mpoly_poly_impl_aux2_def + by (rule mpoly_eqI) + (simp add: coeff_coeff_mpoly_to_mpoly_poly coeff_MPoly + lookup0_fmap_of_list_mpoly_to_mpoly_poly_impl_aux1) + +definition mpoly_to_mpoly_poly_impl :: "nat \ 'a :: comm_ring_1 mpoly \ 'a mpoly list" where + "mpoly_to_mpoly_poly_impl x p = (if p = 0 then [] else + map (mpoly_to_mpoly_poly_impl_aux2 x p) [0.. p = 0" +proof - + interpret transfer_mpoly_to_mpoly_poly x . + define p' where "p' = mpoly_to_mpoly_poly x p" + have [transfer_rule]: "R p' p" + by (auto simp: R_def p'_def) + show ?thesis + unfolding p'_def [symmetric] by transfer_prover +qed + +lemma mpoly_to_mpoly_poly_code [code]: + "Polynomial.coeffs (mpoly_to_mpoly_poly x p) = mpoly_to_mpoly_poly_impl x p" + by (intro nth_equalityI) + (auto simp: mpoly_to_mpoly_poly_impl_def length_coeffs_degree + mpoly_to_mpoly_poly_impl_aux2_def coeffs_nth simp del: upt_Suc) + +value "mpoly_to_mpoly_poly 0 (Var 0 ^ 2 + Var 0 * Var 1 + Var 1 ^ 2 :: int mpoly)" + +value "Rings.divide (Var 0 ^ 2 * Var 1 + Var 0 * Var 1 ^ 2 :: int mpoly) (Var 1)" + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Multivariate_Resultant.thy b/thys/Factor_Algebraic_Polynomial/Multivariate_Resultant.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Multivariate_Resultant.thy @@ -0,0 +1,130 @@ +subsection \Resultants of Multivariate Polynomials\ + +text \We utilize the conversion of multivariate polynomials into univariate polynomials + for the definition of the resultant of multivariate polynomials via + the resultant for univariate polynomials. In this way, we can use the algorithm + to efficiently compute resultants for the multivariate case.\ + +theory Multivariate_Resultant + imports + Poly_Connection + Algebraic_Numbers.Resultant + Subresultants.Subresultant + MPoly_Divide_Code + MPoly_Container +begin + +hide_const (open) + MPoly_Type.degree + MPoly_Type.coeff + Symmetric_Polynomials.lead_coeff + +lemma det_sylvester_matrix_higher_degree: + "det (sylvester_mat_sub (degree f + n) (degree g) f g) + = det (sylvester_mat_sub (degree f) (degree g) f g) * (lead_coeff g * (-1)^(degree g))^n" +proof (induct n) + case (Suc n) + let ?A = "sylvester_mat_sub (degree f + Suc n) (degree g) f g" + let ?d = "degree f + Suc n + degree g" + define h where "h i = ?A $$ (i,0) * cofactor ?A i 0" for i + have mult_left_zero: "x = 0 \ x * y = 0" for x y :: 'a by auto + have "det ?A = (\i = sum h ({degree g} \ ({.. = sum h {degree g} + sum h ({.. = lead_coeff g * cofactor ?A (degree g) 0" unfolding h_def + by (rule arg_cong[of _ _ "\ x. x * _"], simp add: sylvester_mat_sub_def) + also have "cofactor ?A (degree g) 0 = (-1)^(degree g) * det (sylvester_mat_sub (degree f + n) (degree g) f g)" + unfolding cofactor_def + proof (intro arg_cong2[of _ _ _ _ "\ x y. (-1)^x * det y"], force) + show "mat_delete ?A (degree g) 0 = sylvester_mat_sub (degree f + n) (degree g) f g" + unfolding sylvester_mat_sub_def + by (intro eq_matI, auto simp: mat_delete_def coeff_eq_0) + qed + finally show ?case unfolding Suc by simp +qed simp + +text \The conversion of multivariate into univariate polynomials permits us to define resultants in the multivariate + setting. Since in our application one of the polynomials is already univariate, we use a non-symmetric definition + where only one of the input polynomials is multivariate.\ +definition resultant_mpoly_poly :: "nat \ 'a :: comm_ring_1 mpoly \ 'a poly \ 'a mpoly" where + "resultant_mpoly_poly x p q = resultant (mpoly_to_mpoly_poly x p) (map_poly Const q)" + +text \This lemma tells us that there is only a minor difference between computing the multivariate resultant and then + plugging in values, or first inserting values and then evaluate the univariate resultant.\ +lemma insertion_resultant_mpoly_poly: "insertion \ (resultant_mpoly_poly x p q) = resultant (partial_insertion \ x p) q * + (lead_coeff q * (-1)^ degree q)^(degree (mpoly_to_mpoly_poly x p) - degree (partial_insertion \ x p))" +proof - + let ?pa = "partial_insertion \ x" + let ?a = "insertion \" + let ?q = "map_poly Const q" + let ?m = "mpoly_to_mpoly_poly x" + interpret a: comm_ring_hom ?a by (rule comm_ring_hom_insertion) + define m where "m = degree (?m p) - degree (?pa p)" + from degree_partial_insertion_le_mpoly[of \ x p] have deg: "degree (?m p) = degree (?pa p) + m" unfolding m_def by simp + define k where "k = degree (?pa p) + m" + define l where "l = degree q" + have "resultant (?pa p) q = det (sylvester_mat_sub (degree (?pa p)) (degree q) (?pa p) q)" + unfolding resultant_def sylvester_mat_def by simp + have "?a (resultant_mpoly_poly x p q) = ?a (det (sylvester_mat_sub (degree (?pa p) + m) (degree q) (?m p) ?q))" + unfolding resultant_mpoly_poly_def resultant_def sylvester_mat_def degree_map_poly_Const deg .. + also have "\ = + det (a.mat_hom (sylvester_mat_sub (degree (?pa p) + m) (degree q) (?m p) ?q))" + unfolding a.hom_det .. + also have "a.mat_hom (sylvester_mat_sub (degree (?pa p) + m) (degree q) (?m p) ?q) + = sylvester_mat_sub (degree (?pa p) + m) (degree q) (?pa p) q" + unfolding k_def[symmetric] l_def[symmetric] + by (intro eq_matI, auto simp: sylvester_mat_sub_def coeff_map_poly) + also have "det \ = det (sylvester_mat_sub (degree (?pa p)) (degree q) (?pa p) q) * (lead_coeff q * (- 1) ^ degree q) ^ m" + by (subst det_sylvester_matrix_higher_degree, simp) + also have "det (sylvester_mat_sub (degree (?pa p)) (degree q) (?pa p) q) = resultant (?pa p) q" + unfolding resultant_def sylvester_mat_def by simp + finally show ?thesis unfolding m_def by auto +qed + +lemma insertion_resultant_mpoly_poly_zero: fixes q :: "'a :: idom poly" + assumes q: "q \ 0" + shows "insertion \ (resultant_mpoly_poly x p q) = 0 \ resultant (partial_insertion \ x p) q = 0" + unfolding insertion_resultant_mpoly_poly using q by auto + +lemma vars_resultant: "vars (resultant p q) \ \ (vars ` (range (coeff p) \ range (coeff q)))" + unfolding resultant_def det_def sylvester_mat_def sylvester_mat_sub_def + apply simp + apply (rule order.trans[OF vars_setsum]) + subgoal using finite_permutations by blast + apply (rule UN_least) + apply (rule order.trans[OF vars_mult]) + apply simp + apply (rule order.trans[OF vars_prod]) + apply (rule UN_least) + by auto + +text \By taking the resultant, one variable is deleted.\ +lemma vars_resultant_mpoly_poly: "vars (resultant_mpoly_poly x p q) \ vars p - {x}" +proof + fix y + assume "y \ vars (resultant_mpoly_poly x p q)" + from set_mp[OF vars_resultant this[unfolded resultant_mpoly_poly_def]] obtain i + where "y \ vars (coeff (mpoly_to_mpoly_poly x p) i) \ y \ vars (coeff (map_poly Const q) i)" + by auto + moreover have "vars (coeff (map_poly Const q) i) = {}" + by (subst coeff_map_poly, auto) + ultimately have "y \ vars (coeff (mpoly_to_mpoly_poly x p) i)" by auto + thus "y \ More_MPoly_Type.vars p - {x}" using vars_coeff_mpoly_to_mpoly_poly by blast +qed + +text \For resultants, we manually have to select the implementation that + works on integral domains, because there is no factorial ring instance for @{typ "int mpoly"}.\ + +lemma resultant_mpoly_poly_code[code]: + "resultant_mpoly_poly x p q = resultant_impl_basic (mpoly_to_mpoly_poly x p) (map_poly Const q)" + unfolding resultant_mpoly_poly_def div_exp_basic.resultant_impl by simp + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Poly_Connection.thy b/thys/Factor_Algebraic_Polynomial/Poly_Connection.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Poly_Connection.thy @@ -0,0 +1,695 @@ +section \Resultants and Multivariate Polynomials\ + +subsection \Connecting Univariate and Multivariate Polynomials\ + +text \We define a conversion of multivariate polynomials into univariate polynomials + w.r.t.\ a fixed variable $x$ and multivariate polynomials as coefficients.\ + +theory Poly_Connection + imports + Polynomials.MPoly_Type_Univariate + Jordan_Normal_Form.Missing_Misc + Polynomial_Interpolation.Ring_Hom_Poly + Hermite_Lindemann.More_Multivariate_Polynomial_HLW + Polynomials.MPoly_Type_Class +begin + +lemma mpoly_is_unitE: + fixes p :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} mpoly" + assumes "p dvd 1" + obtains c where "p = Const c" "c dvd 1" +proof - + obtain r where r: "p * r = 1" + using assms by auto + from r have [simp]: "p \ 0" "r \ 0" + by auto + have "0 = lead_monom (1 :: 'a mpoly)" + by simp + also have "1 = p * r" + using r by simp + also have "lead_monom (p * r) = lead_monom p + lead_monom r" + by (intro lead_monom_mult) auto + finally have "lead_monom p = 0" + by simp + hence "vars p = {}" + by (simp add: lead_monom_eq_0_iff) + hence *: "p = Const (lead_coeff p)" + by (auto simp: vars_empty_iff) + + have "1 = lead_coeff (1 :: 'a mpoly)" + by simp + also have "1 = p * r" + using r by simp + also have "lead_coeff (p * r) = lead_coeff p * lead_coeff r" + by (intro lead_coeff_mult) auto + finally have "lead_coeff p dvd 1" + using dvdI by blast + with * show ?thesis using that + by blast +qed + +lemma Const_eq_Const_iff [simp]: + "Const c = Const c' \ c = c'" + by (metis lead_coeff_Const) + +lemma is_unit_ConstI [intro]: "c dvd 1 \ Const c dvd 1" + by (metis dvd_def mpoly_Const_1 mpoly_Const_mult) + +lemma is_unit_Const_iff: + fixes c :: "'a :: {comm_semiring_1, semiring_no_zero_divisors}" + shows "Const c dvd 1 \ c dvd 1" +proof + assume "Const c dvd 1" + thus "c dvd 1" + by (auto elim!: mpoly_is_unitE) +qed auto + +lemma vars_emptyE: "vars p = {} \ (\c. p = Const c \ P) \ P" + by (auto simp: vars_empty_iff) + +lemma degree_geI: + assumes "MPoly_Type.coeff p m \ 0" + shows "MPoly_Type.degree p i \ Poly_Mapping.lookup m i" +proof - + have "lookup m i \ Max (insert 0 ((\m. lookup m i) ` keys (mapping_of p)))" + proof (rule Max.coboundedI) + show "lookup m i \ insert 0 ((\m. lookup m i) ` keys (mapping_of p))" + using assms by (auto simp: coeff_keys) + qed auto + thus ?thesis unfolding MPoly_Type.degree_def by auto +qed + +lemma monom_of_degree_exists: + assumes "p \ 0" + obtains m where "MPoly_Type.coeff p m \ 0" "Poly_Mapping.lookup m i = MPoly_Type.degree p i" +proof (cases "MPoly_Type.degree p i = 0") + case False + have "MPoly_Type.degree p i = Max (insert 0 ((\m. lookup m i) ` keys (mapping_of p)))" + by (simp add: MPoly_Type.degree_def) + also have "\ \ insert 0 ((\m. lookup m i) ` keys (mapping_of p))" + by (rule Max_in) auto + finally show ?thesis + using False that by (auto simp: coeff_keys) +next + case [simp]: True + from assms obtain m where m: "MPoly_Type.coeff p m \ 0" + using coeff_all_0 by blast + show ?thesis using degree_geI[of p m i] m + by (intro that[of m]) auto +qed + +lemma degree_leI: + assumes "\m. Poly_Mapping.lookup m i > n \ MPoly_Type.coeff p m = 0" + shows "MPoly_Type.degree p i \ n" +proof (cases "p = 0") + case False + obtain m where m: "MPoly_Type.coeff p m \ 0" "Poly_Mapping.lookup m i = MPoly_Type.degree p i" + using monom_of_degree_exists False by blast + with assms show ?thesis + by force +qed auto + +lemma coeff_gt_degree_eq_0: + assumes "Poly_Mapping.lookup m i > MPoly_Type.degree p i" + shows "MPoly_Type.coeff p m = 0" + using assms degree_geI leD by blast + +lemma vars_altdef: "vars p = (\m\{m. MPoly_Type.coeff p m \ 0}. keys m)" + unfolding vars_def + by (intro arg_cong[where f = "\"] image_cong refl) (simp flip: coeff_keys) + +lemma degree_pos_iff: "MPoly_Type.degree p x > 0 \ x \ vars p" +proof + assume "MPoly_Type.degree p x > 0" + hence "p \ 0" by auto + then obtain m where m: "lookup m x = MPoly_Type.degree p x" "MPoly_Type.coeff p m \ 0" + using monom_of_degree_exists[of p x] by metis + from m and \MPoly_Type.degree p x > 0\ have "x \ keys m" + by (simp add: in_keys_iff) + with m show "x \ vars p" + by (auto simp: vars_altdef) +next + assume "x \ vars p" + then obtain m where m: "x \ keys m" "MPoly_Type.coeff p m \ 0" + by (auto simp: vars_altdef) + have "0 < lookup m x" + using m by (auto simp: in_keys_iff) + also from m have "\ \ MPoly_Type.degree p x" + by (intro degree_geI) auto + finally show "MPoly_Type.degree p x > 0" . +qed + +lemma degree_eq_0_iff: "MPoly_Type.degree p x = 0 \ x \ vars p" + using degree_pos_iff[of p x] by auto + +lemma MPoly_Type_monom_zero[simp]: "MPoly_Type.monom m 0 = 0" + by (simp add: More_MPoly_Type.coeff_monom coeff_all_0) + +lemma vars_monom_keys': "vars (MPoly_Type.monom m c) = (if c = 0 then {} else keys m)" + by (cases "c = 0") (auto simp: vars_monom_keys) + +lemma Const_eq_0_iff [simp]: "Const c = 0 \ c = 0" + by (metis lead_coeff_Const mpoly_Const_0) + +lemma monom_remove_key: "MPoly_Type.monom m (a :: 'a :: semiring_1) = + MPoly_Type.monom (remove_key x m) a * MPoly_Type.monom (Poly_Mapping.single x (lookup m x)) 1" + unfolding MPoly_Type.mult_monom + by (rule arg_cong2[of _ _ _ _ MPoly_Type.monom], auto simp: remove_key_sum) + +lemma MPoly_Type_monom_0_iff[simp]: "MPoly_Type.monom m x = 0 \ x = 0" + by (metis (full_types) MPoly_Type_monom_zero More_MPoly_Type.coeff_monom when_def) + +lemma vars_signof[simp]: "vars (signof x) = {}" + by (simp add: sign_def) + +lemma prod_mset_Const: "prod_mset (image_mset Const A) = Const (prod_mset A)" + by (induction A) (auto simp: mpoly_Const_mult) + +lemma Const_eq_product_iff: + fixes c :: "'a :: idom" + assumes "c \ 0" + shows "Const c = a * b \ (\a' b'. a = Const a' \ b = Const b' \ c = a' * b')" +proof + assume *: "Const c = a * b" + have "lead_monom (a * b) = 0" + by (auto simp flip: *) + hence "lead_monom a = 0 \ lead_monom b = 0" + by (subst (asm) lead_monom_mult) (use assms * in auto) + hence "vars a = {}" "vars b = {}" + by (auto simp: lead_monom_eq_0_iff) + then obtain a' b' where "a = Const a'" "b = Const b'" + by (auto simp: vars_empty_iff) + with * show "(\a' b'. a = Const a' \ b = Const b' \ c = a' * b')" + by (auto simp flip: mpoly_Const_mult) +qed (auto simp: mpoly_Const_mult) + +lemma irreducible_Const_iff [simp]: + "irreducible (Const (c :: 'a :: idom)) \ irreducible c" +proof + assume *: "irreducible (Const c)" + show "irreducible c" + proof (rule irreducibleI) + fix a b assume "c = a * b" + hence "Const c = Const a * Const b" + by (simp add: mpoly_Const_mult) + with * have "Const a dvd 1 \ Const b dvd 1" + by blast + thus "a dvd 1 \ b dvd 1" + by (meson is_unit_Const_iff) + qed (use * in \auto simp: irreducible_def\) +next + assume *: "irreducible c" + have [simp]: "c \ 0" + using * by auto + show "irreducible (Const c)" + proof (rule irreducibleI) + fix a b assume "Const c = a * b" + then obtain a' b' where [simp]: "a = Const a'" "b = Const b'" and "c = a' * b'" + by (auto simp: Const_eq_product_iff) + hence "a' dvd 1 \ b' dvd 1" + using * by blast + thus "a dvd 1 \ b dvd 1" + by auto + qed (use * in \auto simp: irreducible_def is_unit_Const_iff\) +qed + +lemma Const_dvd_Const_iff [simp]: "Const a dvd Const b \ a dvd b" +proof + assume "a dvd b" + then obtain c where "b = a * c" + by auto + hence "Const b = Const a * Const c" + by (auto simp: mpoly_Const_mult) + thus "Const a dvd Const b" + by simp +next + assume "Const a dvd Const b" + then obtain p where p: "Const b = Const a * p" + by auto + have "MPoly_Type.coeff (Const b) 0 = MPoly_Type.coeff (Const a * p) 0" + using p by simp + also have "\ = MPoly_Type.coeff (Const a) 0 * MPoly_Type.coeff p 0" + using mpoly_coeff_times_0 by blast + finally show "a dvd b" + by (simp add: mpoly_coeff_Const) +qed + + +text \The lemmas above should be moved into the right theories. The part below is on the new +connection between multivariate polynomials and univariate polynomials.\ + +text \The imported theories only allow a conversion from one-variable mpoly's to poly and vice-versa. + However, we require a conversion from arbitrary mpoly's into poly's with mpolys as coefficients.\ + +(* converts a multi-variate polynomial into a univariate polynomial with multivariate coefficients *) +definition mpoly_to_mpoly_poly :: "nat \ 'a :: comm_ring_1 mpoly \ 'a mpoly poly" where + "mpoly_to_mpoly_poly x p = (\m . + Polynomial.monom (MPoly_Type.monom (remove_key x m) (MPoly_Type.coeff p m)) (lookup m x))" + +lemma mpoly_to_mpoly_poly_add [simp]: + "mpoly_to_mpoly_poly x (p + q) = mpoly_to_mpoly_poly x p + mpoly_to_mpoly_poly x q" + unfolding mpoly_to_mpoly_poly_def More_MPoly_Type.coeff_add[symmetric] MPoly_Type.monom_add add_monom[symmetric] + by (rule Sum_any.distrib) auto + +lemma mpoly_to_mpoly_poly_monom: "mpoly_to_mpoly_poly x (MPoly_Type.monom m a) = Polynomial.monom (MPoly_Type.monom (remove_key x m) a) (lookup m x)" +proof - + have "mpoly_to_mpoly_poly x (MPoly_Type.monom m a) = + (\ m'. Polynomial.monom (MPoly_Type.monom (remove_key x m') a) (lookup m' x) when m' = m)" + unfolding mpoly_to_mpoly_poly_def + by (intro Sum_any.cong, auto simp: when_def More_MPoly_Type.coeff_monom) + also have "\ = Polynomial.monom (MPoly_Type.monom (remove_key x m) a) (lookup m x)" + unfolding Sum_any_when_equal .. + finally show ?thesis . +qed + +lemma remove_key_transfer [transfer_rule]: + "rel_fun (=) (rel_fun (pcr_poly_mapping (=) (=)) (pcr_poly_mapping (=) (=))) + (\k0 f k. f k when k \ k0) remove_key" + unfolding pcr_poly_mapping_def cr_poly_mapping_def OO_def + by (auto simp: rel_fun_def remove_key_lookup) + +lemma remove_key_0 [simp]: "remove_key x 0 = 0" + by transfer auto + +lemma remove_key_single' [simp]: + "x \ y \ remove_key x (Poly_Mapping.single y n) = Poly_Mapping.single y n" + by transfer (auto simp: when_def fun_eq_iff) + + +lemma poly_coeff_Sum_any: + assumes "finite {x. f x \ 0}" + shows "poly.coeff (Sum_any f) n = Sum_any (\x. poly.coeff (f x) n)" +proof - + have "Sum_any f = (\x | f x \ 0. f x)" + by (rule Sum_any.expand_set) + also have "poly.coeff \ n = (\x | f x \ 0. poly.coeff (f x) n)" + by (simp add: Polynomial.coeff_sum) + also have "\ = Sum_any (\x. poly.coeff (f x) n)" + by (rule Sum_any.expand_superset [symmetric]) (use assms in auto) + finally show ?thesis . +qed + +lemma coeff_coeff_mpoly_to_mpoly_poly: + "MPoly_Type.coeff (poly.coeff (mpoly_to_mpoly_poly x p) n) m = + (MPoly_Type.coeff p (m + Poly_Mapping.single x n) when lookup m x = 0)" +proof - + have "MPoly_Type.coeff (poly.coeff (mpoly_to_mpoly_poly x p) n) m = + MPoly_Type.coeff (\a. MPoly_Type.monom (remove_key x a) (MPoly_Type.coeff p a) when lookup a x = n) m" + unfolding mpoly_to_mpoly_poly_def by (subst poly_coeff_Sum_any) (auto simp: when_def) + also have "\ = (\xa. MPoly_Type.coeff (MPoly_Type.monom (remove_key x xa) (MPoly_Type.coeff p xa)) m when lookup xa x = n)" + by (subst coeff_Sum_any, force) (auto simp: when_def intro!: Sum_any.cong) + also have "\ = (\a. MPoly_Type.coeff p a when lookup a x = n \ m = remove_key x a)" + by (intro Sum_any.cong) (simp add: More_MPoly_Type.coeff_monom when_def) + also have "(\a. lookup a x = n \ m = remove_key x a) = + (\a. lookup m x = 0 \ a = m + Poly_Mapping.single x n)" + by (rule ext, transfer) (auto simp: fun_eq_iff when_def) + also have "(\a. MPoly_Type.coeff p a when \ a) = + (\a. MPoly_Type.coeff p a when lookup m x = 0 when a = m + Poly_Mapping.single x n)" + by (intro Sum_any.cong) (auto simp: when_def) + also have "\ = (MPoly_Type.coeff p (m + Poly_Mapping.single x n) when lookup m x = 0)" + by (rule Sum_any_when_equal) + finally show ?thesis . +qed + +lemma mpoly_to_mpoly_poly_Const [simp]: + "mpoly_to_mpoly_poly x (Const c) = [:Const c:]" +proof - + have "mpoly_to_mpoly_poly x (Const c) = + (\m. Polynomial.monom (MPoly_Type.monom (remove_key x m) + (MPoly_Type.coeff (Const c) m)) (lookup m x) when m = 0)" + unfolding mpoly_to_mpoly_poly_def + by (intro Sum_any.cong) (auto simp: when_def mpoly_coeff_Const) + also have "\ = [:Const c:]" + by (subst Sum_any_when_equal) + (auto simp: mpoly_coeff_Const monom_altdef simp flip: Const_conv_monom) + finally show ?thesis . +qed + +lemma mpoly_to_mpoly_poly_Var: + "mpoly_to_mpoly_poly x (Var y) = (if x = y then [:0, 1:] else [:Var y:])" +proof - + have "mpoly_to_mpoly_poly x (Var y) = + (\a. Polynomial.monom (MPoly_Type.monom (remove_key x a) 1) (lookup a x) + when a = Poly_Mapping.single y 1)" + unfolding mpoly_to_mpoly_poly_def by (intro Sum_any.cong) (auto simp: when_def coeff_Var) + also have "\ = (if x = y then [:0, 1:] else [:Var y:])" + by (auto simp: Polynomial.monom_altdef lookup_single Var_altdef) + finally show ?thesis . +qed + +lemma mpoly_to_mpoly_poly_Var_this [simp]: + "mpoly_to_mpoly_poly x (Var x) = [:0, 1:]" + "x \ y \ mpoly_to_mpoly_poly x (Var y) = [:Var y:]" + by (simp_all add: mpoly_to_mpoly_poly_Var) + +lemma mpoly_to_mpoly_poly_uminus [simp]: + "mpoly_to_mpoly_poly x (-p) = -mpoly_to_mpoly_poly x p" + unfolding mpoly_to_mpoly_poly_def + by (auto simp: monom_uminus Sum_any_uminus simp flip: minus_monom) + +lemma mpoly_to_mpoly_poly_diff [simp]: + "mpoly_to_mpoly_poly x (p - q) = mpoly_to_mpoly_poly x p - mpoly_to_mpoly_poly x q" + by (subst diff_conv_add_uminus, subst mpoly_to_mpoly_poly_add) auto + +lemma mpoly_to_mpoly_poly_0 [simp]: + "mpoly_to_mpoly_poly x 0 = 0" + unfolding mpoly_Const_0 [symmetric] mpoly_to_mpoly_poly_Const by simp + +lemma mpoly_to_mpoly_poly_1 [simp]: + "mpoly_to_mpoly_poly x 1 = 1" + unfolding mpoly_Const_1 [symmetric] mpoly_to_mpoly_poly_Const by simp + +lemma mpoly_to_mpoly_poly_of_nat [simp]: + "mpoly_to_mpoly_poly x (of_nat n) = of_nat n" + unfolding of_nat_mpoly_eq mpoly_to_mpoly_poly_Const of_nat_poly .. + +lemma mpoly_to_mpoly_poly_of_int [simp]: + "mpoly_to_mpoly_poly x (of_int n) = of_int n" + unfolding of_nat_mpoly_eq mpoly_to_mpoly_poly_Const of_nat_poly by (cases n) auto + +lemma mpoly_to_mpoly_poly_numeral [simp]: + "mpoly_to_mpoly_poly x (numeral n) = numeral n" + using mpoly_to_mpoly_poly_of_nat[of x "numeral n"] by (simp del: mpoly_to_mpoly_poly_of_nat) + +lemma coeff_monom_mult': + "MPoly_Type.coeff (MPoly_Type.monom m a * q) m' = + (a * MPoly_Type.coeff q (m' - m) when lookup m' \ lookup m)" +proof (cases "lookup m' \ lookup m") + case True + have "a * MPoly_Type.coeff q (m' - m) = MPoly_Type.coeff (MPoly_Type.monom m a * q) (m + (m' - m))" + by (rule More_MPoly_Type.coeff_monom_mult [symmetric]) + also have "m + (m' - m) = m'" + using True by transfer (auto simp: le_fun_def) + finally show ?thesis + using True by (simp add: when_def) +next + case False + have "MPoly_Type.coeff (MPoly_Type.monom m a * q) m' = + (\m1. a * (\m2. MPoly_Type.coeff q m2 when m' = m1 + m2) when m1 = m)" + unfolding coeff_mpoly_times prod_fun_def + by (intro Sum_any.cong) (auto simp: More_MPoly_Type.coeff_monom when_def) + also have "\ = a * (\m2. MPoly_Type.coeff q m2 when m' = m + m2)" + by (subst Sum_any_when_equal) auto + also have "(\m2. m' = m + m2) = (\m2. False)" + by (rule ext) (use False in \transfer, auto simp: le_fun_def\) + finally show ?thesis + using False by simp +qed + +lemma mpoly_to_mpoly_poly_mult_monom: + "mpoly_to_mpoly_poly x (MPoly_Type.monom m a * q) = + Polynomial.monom (MPoly_Type.monom (remove_key x m) a) (lookup m x) * mpoly_to_mpoly_poly x q" + (is "?lhs = ?rhs") +proof (rule poly_eqI, rule mpoly_eqI) + fix n :: nat and mon :: "nat \\<^sub>0 nat" + have "MPoly_Type.coeff (poly.coeff ?lhs n) mon = + (a * MPoly_Type.coeff q (mon + Poly_Mapping.single x n - m) + when lookup m \ lookup (mon + Poly_Mapping.single x n) \ lookup mon x = 0)" + by (simp add: coeff_coeff_mpoly_to_mpoly_poly coeff_monom_mult' when_def) + have "MPoly_Type.coeff (poly.coeff ?rhs n) mon = + (a * MPoly_Type.coeff q (mon - remove_key x m + Poly_Mapping.single x (n - lookup m x)) + when lookup (remove_key x m) \ lookup mon \ lookup m x \ n \ lookup mon x = 0)" + by (simp add: coeff_coeff_mpoly_to_mpoly_poly coeff_monom_mult' lookup_minus_fun + remove_key_lookup Missing_Polynomial.coeff_monom_mult when_def) + also have "lookup (remove_key x m) \ lookup mon \ lookup m x \ n \ lookup mon x = 0 \ + lookup m \ lookup (mon + Poly_Mapping.single x n) \ lookup mon x = 0" (is "_ = ?P") + by transfer (auto simp: when_def le_fun_def) + also have "mon - remove_key x m + Poly_Mapping.single x (n - lookup m x) = mon + Poly_Mapping.single x n - m" if ?P + using that by transfer (auto simp: fun_eq_iff when_def) + hence "(a * MPoly_Type.coeff q (mon - remove_key x m + Poly_Mapping.single x (n - lookup m x)) when ?P) = + (a * MPoly_Type.coeff q \ when ?P)" + by (intro when_cong) auto + also have "\ = MPoly_Type.coeff (poly.coeff ?lhs n) mon" + by (simp add: coeff_coeff_mpoly_to_mpoly_poly coeff_monom_mult' when_def) + finally show "MPoly_Type.coeff (poly.coeff ?lhs n) mon = MPoly_Type.coeff (poly.coeff ?rhs n) mon" .. +qed + +lemma mpoly_to_mpoly_poly_mult [simp]: + "mpoly_to_mpoly_poly x (p * q) = mpoly_to_mpoly_poly x p * mpoly_to_mpoly_poly x q" + by (induction p arbitrary: q rule: mpoly_induct) + (simp_all add: mpoly_to_mpoly_poly_monom mpoly_to_mpoly_poly_mult_monom ring_distribs) + +lemma coeff_mpoly_to_mpoly_poly: + "Polynomial.coeff (mpoly_to_mpoly_poly x p) n = + Sum_any (\m. MPoly_Type.monom (remove_key x m) (MPoly_Type.coeff p m) when Poly_Mapping.lookup m x = n)" + unfolding mpoly_to_mpoly_poly_def by (subst poly_coeff_Sum_any) (auto simp: when_def) + +lemma mpoly_coeff_to_mpoly_poly_coeff: + "MPoly_Type.coeff p m = MPoly_Type.coeff (poly.coeff (mpoly_to_mpoly_poly x p) (lookup m x)) (remove_key x m)" +proof - + have "MPoly_Type.coeff (poly.coeff (mpoly_to_mpoly_poly x p) (lookup m x)) (remove_key x m) = + (\xa. MPoly_Type.coeff (MPoly_Type.monom (remove_key x xa) (MPoly_Type.coeff p xa) when + lookup xa x = lookup m x) (remove_key x m))" + by (subst coeff_mpoly_to_mpoly_poly, subst coeff_Sum_any) auto + also have "\ = (\xa. MPoly_Type.coeff (MPoly_Type.monom (remove_key x xa) (MPoly_Type.coeff p xa)) (remove_key x m) + when lookup xa x = lookup m x)" + by (intro Sum_any.cong) (auto simp: when_def) + also have "\ = (\xa. MPoly_Type.coeff p xa when remove_key x m = remove_key x xa \ lookup xa x = lookup m x)" + by (intro Sum_any.cong) (auto simp: More_MPoly_Type.coeff_monom when_def) + also have "(\xa. remove_key x m = remove_key x xa \ lookup xa x = lookup m x) = (\xa. xa = m)" + using remove_key_sum by metis + also have "(\xa. MPoly_Type.coeff p xa when xa = m) = MPoly_Type.coeff p m" + by simp + finally show ?thesis .. +qed + +lemma degree_mpoly_to_mpoly_poly [simp]: + "Polynomial.degree (mpoly_to_mpoly_poly x p) = MPoly_Type.degree p x" +proof (rule antisym) + show "Polynomial.degree (mpoly_to_mpoly_poly x p) \ MPoly_Type.degree p x" + proof (intro Polynomial.degree_le allI impI) + fix i assume i: "i > MPoly_Type.degree p x" + have "poly.coeff (mpoly_to_mpoly_poly x p) i = + (\m. 0 when lookup m x = i)" + unfolding coeff_mpoly_to_mpoly_poly using i + by (intro Sum_any.cong when_cong refl) (auto simp: coeff_gt_degree_eq_0) + also have "\ = 0" + by simp + finally show "poly.coeff (mpoly_to_mpoly_poly x p) i = 0" . + qed +next + show "Polynomial.degree (mpoly_to_mpoly_poly x p) \ MPoly_Type.degree p x" + proof (cases "p = 0") + case False + then obtain m where m: "MPoly_Type.coeff p m \ 0" "lookup m x = MPoly_Type.degree p x" + using monom_of_degree_exists by blast + show "Polynomial.degree (mpoly_to_mpoly_poly x p) \ MPoly_Type.degree p x" + proof (rule Polynomial.le_degree) + have "0 \ MPoly_Type.coeff p m" + using m by auto + also have "MPoly_Type.coeff p m = MPoly_Type.coeff (poly.coeff (mpoly_to_mpoly_poly x p) (lookup m x)) (remove_key x m)" + by (rule mpoly_coeff_to_mpoly_poly_coeff) + finally show "poly.coeff (mpoly_to_mpoly_poly x p) (MPoly_Type.degree p x) \ 0" + using m by auto + qed + qed auto +qed + +text \The upcoming lemma is similar to @{thm reduce_nested_mpoly_extract_var}.\ +lemma poly_mpoly_to_mpoly_poly: + "poly (mpoly_to_mpoly_poly x p) (Var x) = p" +proof (induct p rule: mpoly_induct) + case (monom m a) + show ?case unfolding mpoly_to_mpoly_poly_monom poly_monom + by (transfer, simp add: Var\<^sub>0_power mult_single remove_key_sum) +next + case (sum p1 p2 m a) + then show ?case by (simp add: mpoly_to_mpoly_poly_add) +qed + +lemma mpoly_to_mpoly_poly_eq_iff [simp]: + "mpoly_to_mpoly_poly x p = mpoly_to_mpoly_poly x q \ p = q" +proof + assume "mpoly_to_mpoly_poly x p = mpoly_to_mpoly_poly x q" + hence "poly (mpoly_to_mpoly_poly x p) (Var x) = + poly (mpoly_to_mpoly_poly x q) (Var x)" + by simp + thus "p = q" + by (auto simp: poly_mpoly_to_mpoly_poly) +qed auto + +text \Evaluation, i.e., insertion of concrete values is identical\ +lemma insertion_mpoly_to_mpoly_poly: assumes "\ y. y \ x \ \ y = \ y" + shows "poly (map_poly (insertion \) (mpoly_to_mpoly_poly x p)) (\ x) = insertion \ p" +proof (induct p rule: mpoly_induct) + case (monom m a) + let ?rkm = "remove_key x m" + have to_alpha: "insertion \ (MPoly_Type.monom ?rkm a) = insertion \ (MPoly_Type.monom ?rkm a)" + by (rule insertion_irrelevant_vars, rule assms, insert vars_monom_subset[of ?rkm a], auto simp: remove_key_keys[symmetric]) + have main: "insertion \ (MPoly_Type.monom ?rkm a) * \ x ^ lookup m x = insertion \ (MPoly_Type.monom m a)" + unfolding monom_remove_key[of m a x] insertion_mult + by (metis insertion_single mult.left_neutral) + show ?case using main to_alpha + by (simp add: mpoly_to_mpoly_poly_monom map_poly_monom poly_monom) +next + case (sum p1 p2 m a) + then show ?case by (simp add: mpoly_to_mpoly_poly_add insertion_add map_poly_add) +qed + +lemma mpoly_to_mpoly_poly_dvd_iff [simp]: + "mpoly_to_mpoly_poly x p dvd mpoly_to_mpoly_poly x q \ p dvd q" +proof + assume "mpoly_to_mpoly_poly x p dvd mpoly_to_mpoly_poly x q" + hence "poly (mpoly_to_mpoly_poly x p) (Var x) dvd poly (mpoly_to_mpoly_poly x q) (Var x)" + by (intro poly_hom.hom_dvd) + thus "p dvd q" + by (simp add: poly_mpoly_to_mpoly_poly) +qed auto + +lemma vars_coeff_mpoly_to_mpoly_poly: "vars (poly.coeff (mpoly_to_mpoly_poly x p) i) \ vars p - {x}" + unfolding mpoly_to_mpoly_poly_def Sum_any.expand_set Polynomial.coeff_sum More_MPoly_Type.coeff_monom + apply (rule order.trans[OF vars_setsum], force) + apply (rule UN_least, simp) + apply (intro impI order.trans[OF vars_monom_subset]) + by (simp add: remove_key_keys[symmetric] Diff_mono SUP_upper2 coeff_keys vars_def) + + +locale transfer_mpoly_to_mpoly_poly = + fixes x :: nat +begin + +definition R :: "'a :: comm_ring_1 mpoly poly \ 'a mpoly \ bool" where + "R p p' \ p = mpoly_to_mpoly_poly x p'" + +context + includes lifting_syntax +begin + +lemma transfer_0 [transfer_rule]: "R 0 0" + and transfer_1 [transfer_rule]: "R 1 1" + and transfer_Const [transfer_rule]: "R [:Const c:] (Const c)" + and transfer_uminus [transfer_rule]: "(R ===> R) uminus uminus" + and transfer_of_nat [transfer_rule]: "((=) ===> R) of_nat of_nat" + and transfer_of_int [transfer_rule]: "((=) ===> R) of_nat of_nat" + and transfer_numeral [transfer_rule]: "((=) ===> R) of_nat of_nat" + and transfer_add [transfer_rule]: "(R ===> R ===> R) (+) (+)" + and transfer_diff [transfer_rule]: "(R ===> R ===> R) (+) (+)" + and transfer_mult [transfer_rule]: "(R ===> R ===> R) (*) (*)" + and transfer_dvd [transfer_rule]: "(R ===> R ===> (=)) (dvd) (dvd)" + and transfer_monom [transfer_rule]: + "((=) ===> (=) ===> R) + (\m a. Polynomial.monom (MPoly_Type.monom (remove_key x m) a) (lookup m x)) + MPoly_Type.monom" + and transfer_coeff [transfer_rule]: + "(R ===> (=) ===> (=)) + (\p m. MPoly_Type.coeff (poly.coeff p (lookup m x)) (remove_key x m)) + MPoly_Type.coeff" + and transfer_degree [transfer_rule]: + "(R ===> (=)) Polynomial.degree (\p. MPoly_Type.degree p x)" + unfolding R_def + by (auto simp: rel_fun_def mpoly_to_mpoly_poly_monom + simp flip: mpoly_coeff_to_mpoly_poly_coeff) + + +lemma transfer_vars [transfer_rule]: + assumes [transfer_rule]: "R p p'" + shows "(\i. vars (poly.coeff p i)) \ (if Polynomial.degree p = 0 then {} else {x}) = vars p'" + (is "?A \ ?B = _") +proof (intro equalityI) + have "vars p' = vars (poly p (Var x))" + using assms by (simp add: R_def poly_mpoly_to_mpoly_poly) + also have "poly p (Var x) = (\i\Polynomial.degree p. poly.coeff p i * Var x ^ i)" + unfolding poly_altdef .. + also have "vars \ \ (\i. vars (poly.coeff p i) \ (if Polynomial.degree p = 0 then {} else {x}))" + proof (intro order.trans[OF vars_sum] UN_mono order.trans[OF vars_mult] Un_mono) + fix i :: nat + assume i: "i \ {..Polynomial.degree p}" + show "vars (Var x ^ i) \ (if Polynomial.degree p = 0 then {} else {x})" + proof (cases "Polynomial.degree p = 0") + case False + thus ?thesis + by (intro order.trans[OF vars_power]) (auto simp: vars_Var) + qed (use i in auto) + qed auto + finally show "vars p' \ ?A \ ?B" by blast +next + have "?A \ vars p'" + using assms vars_coeff_mpoly_to_mpoly_poly by (auto simp: R_def) + moreover have "?B \ vars p'" + using assms by (auto simp: R_def degree_pos_iff) + ultimately show "?A \ ?B \ vars p'" + by blast +qed + +lemma right_total [transfer_rule]: "right_total R" + unfolding right_total_def +proof safe + fix p' :: "'a mpoly" + show "\p. R p p'" + by (rule exI[of _ "mpoly_to_mpoly_poly x p'"]) (auto simp: R_def) +qed + +lemma bi_unique [transfer_rule]: "bi_unique R" + unfolding bi_unique_def by (auto simp: R_def) + +end + +end + + +lemma mpoly_degree_mult_eq: + fixes p q :: "'a :: idom mpoly" + assumes "p \ 0" "q \ 0" + shows "MPoly_Type.degree (p * q) x = MPoly_Type.degree p x + MPoly_Type.degree q x" +proof - + interpret transfer_mpoly_to_mpoly_poly x . + define deg :: "'a mpoly \ nat" where "deg = (\p. MPoly_Type.degree p x)" + have [transfer_rule]: "rel_fun R (=) Polynomial.degree deg" + using transfer_degree unfolding deg_def . + + have "deg (p * q) = deg p + deg q" + using assms unfolding deg_def [symmetric] + by transfer (simp add: degree_mult_eq) + thus ?thesis + by (simp add: deg_def) +qed + + + +text \Converts a multi-variate polynomial into a univariate polynomial via inserting values for all but one variable\ +definition partial_insertion :: "(nat \ 'a) \ nat \ 'a :: comm_ring_1 mpoly \ 'a poly" where + "partial_insertion \ x p = map_poly (insertion \) (mpoly_to_mpoly_poly x p)" + +lemma comm_ring_hom_insertion: "comm_ring_hom (insertion \)" + by (unfold_locales, auto simp: insertion_add insertion_mult) + + +lemma partial_insertion_add: "partial_insertion \ x (p + q) = partial_insertion \ x p + partial_insertion \ x q" +proof - + interpret i: comm_ring_hom "insertion \" by (rule comm_ring_hom_insertion) + show ?thesis unfolding partial_insertion_def mpoly_to_mpoly_poly_add hom_distribs .. +qed + +lemma partial_insertion_monom: "partial_insertion \ x (MPoly_Type.monom m a) = Polynomial.monom (insertion \ (MPoly_Type.monom (remove_key x m) a)) (lookup m x)" + unfolding partial_insertion_def mpoly_to_mpoly_poly_monom + by (subst map_poly_monom, auto) + +text \Partial insertion + insertion of last value is identical to (full) insertion\ +lemma insertion_partial_insertion: assumes "\ y. y \ x \ \ y = \ y" + shows "poly (partial_insertion \ x p) (\ x) = insertion \ p" +proof (induct p rule: mpoly_induct) + case (monom m a) + let ?rkm = "remove_key x m" + have to_alpha: "insertion \ (MPoly_Type.monom ?rkm a) = insertion \ (MPoly_Type.monom ?rkm a)" + by (rule insertion_irrelevant_vars, rule assms, insert vars_monom_subset[of ?rkm a], auto simp: remove_key_keys[symmetric]) + have main: "insertion \ (MPoly_Type.monom ?rkm a) * \ x ^ lookup m x = insertion \ (MPoly_Type.monom m a)" + unfolding monom_remove_key[of m a x] insertion_mult + by (metis insertion_single mult.left_neutral) + show ?case using main to_alpha by (simp add: partial_insertion_monom poly_monom) +next + case (sum p1 p2 m a) + then show ?case by (simp add: partial_insertion_add insertion_add map_poly_add) +qed + +lemma insertion_coeff_mpoly_to_mpoly_poly[simp]: + "insertion \ (coeff (mpoly_to_mpoly_poly x p) k) = coeff (partial_insertion \ x p) k" + unfolding partial_insertion_def + by (subst coeff_map_poly, auto) + +lemma degree_map_poly_Const: "degree (map_poly (Const :: 'a :: semiring_0 \ _) f) = degree f" + by (rule degree_map_poly, auto) + +lemma degree_partial_insertion_le_mpoly: "degree (partial_insertion \ x p) \ degree (mpoly_to_mpoly_poly x p)" + unfolding partial_insertion_def by (rule degree_map_poly_le) + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/ROOT b/thys/Factor_Algebraic_Polynomial/ROOT new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/ROOT @@ -0,0 +1,14 @@ +chapter AFP + +session Factor_Algebraic_Polynomial(AFP) = Algebraic_Numbers + + options [timeout = 600] + sessions + Polynomials + Hermite_Lindemann + theories + Poly_Connection + Multivariate_Resultant + Factor_Real_Poly + document_files + "root.tex" + "root.bib" diff --git a/thys/Factor_Algebraic_Polynomial/Roots_of_Algebraic_Poly.thy b/thys/Factor_Algebraic_Polynomial/Roots_of_Algebraic_Poly.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Roots_of_Algebraic_Poly.thy @@ -0,0 +1,931 @@ +section \Representing Roots of Polynomials with Algebraic Coefficients\ + +text \We provide an algorithm to compute a non-zero integer polynomial $q$ from a polynomial + $p$ with algebraic coefficients such that all roots of $p$ are also roots of $q$. + + In this way, we have a constructive proof that the set of complex algebraic numbers + is algebraically closed.\ + +theory Roots_of_Algebraic_Poly + imports + Algebraic_Numbers.Complex_Algebraic_Numbers + Multivariate_Resultant + Is_Int_To_Int +begin + +subsection \Preliminaries\ + +hide_const (open) up_ring.monom +hide_const (open) MPoly_Type.monom + +lemma map_mpoly_Const: "f 0 = 0 \ map_mpoly f (Const i) = Const (f i)" + by (intro mpoly_eqI, auto simp: coeff_map_mpoly mpoly_coeff_Const) + +lemma map_mpoly_Var: "f 1 = 1 \ map_mpoly (f :: 'b :: zero_neq_one \ _) (Var i) = Var i" + by (intro mpoly_eqI, auto simp: coeff_map_mpoly coeff_Var when_def) + +lemma map_mpoly_monom: "f 0 = 0 \ map_mpoly f (MPoly_Type.monom m a) = (MPoly_Type.monom m (f a))" + by (intro mpoly_eqI, unfold coeff_map_mpoly if_distrib coeff_monom, simp add: when_def) + +lemma remove_key_single': + "remove_key v (Poly_Mapping.single w n) = (if v = w then 0 else Poly_Mapping.single w n)" + by (metis add.right_neutral lookup_single_not_eq remove_key_single remove_key_sum single_zero) + +context comm_monoid_add_hom +begin +lemma hom_Sum_any: assumes fin: "finite {x. f x \ 0}" + shows "hom (Sum_any f) = Sum_any (\ x. hom (f x))" + unfolding Sum_any.expand_set hom_sum + by (rule sum.mono_neutral_right[OF fin], auto) + +lemma comm_monoid_add_hom_mpoly_map: "comm_monoid_add_hom (map_mpoly hom)" + by (unfold_locales; intro mpoly_eqI, auto simp: hom_add) + +lemma map_mpoly_hom_Const: "map_mpoly hom (Const i) = Const (hom i)" + by (rule map_mpoly_Const, simp) + +lemma map_mpoly_hom_monom: "map_mpoly hom (MPoly_Type.monom m a) = MPoly_Type.monom m (hom a)" + by (rule map_mpoly_monom, simp) +end + +context comm_ring_hom +begin +lemma mpoly_to_poly_map_mpoly_hom: "mpoly_to_poly x (map_mpoly hom p) = map_poly hom (mpoly_to_poly x p)" + by (rule poly_eqI, unfold coeff_mpoly_to_poly coeff_map_poly_hom, subst coeff_map_mpoly', auto) + +lemma comm_ring_hom_mpoly_map: "comm_ring_hom (map_mpoly hom)" +proof - + interpret mp: comm_monoid_add_hom "map_mpoly hom" by (rule comm_monoid_add_hom_mpoly_map) + show ?thesis + proof (unfold_locales) + show "map_mpoly hom 1 = 1" + by (intro mpoly_eqI, simp add: MPoly_Type.coeff_def, transfer fixing: hom, transfer fixing: hom, auto simp: when_def) + fix x y + show "map_mpoly hom (x * y) = map_mpoly hom x * map_mpoly hom y" + apply (intro mpoly_eqI) + apply (subst coeff_map_mpoly', force) + apply (unfold coeff_mpoly_times) + apply (subst prod_fun_unfold_prod, blast, blast) + apply (subst prod_fun_unfold_prod, blast, blast) + apply (subst coeff_map_mpoly', force) + apply (subst coeff_map_mpoly', force) + apply (subst hom_Sum_any) + subgoal + proof - + let ?X = "{a. MPoly_Type.coeff x a \ 0}" + let ?Y = "{a. MPoly_Type.coeff y a \ 0}" + have fin: "finite (?X \ ?Y)" by auto + show ?thesis + by (rule finite_subset[OF _ fin], auto) + qed + apply (rule Sum_any.cong) + subgoal for mon pair by (cases pair, auto simp: hom_mult when_def) + done + qed +qed + +lemma mpoly_to_mpoly_poly_map_mpoly_hom: + "mpoly_to_mpoly_poly x (map_mpoly hom p) = map_poly (map_mpoly hom) (mpoly_to_mpoly_poly x p)" +proof - + interpret mp: comm_ring_hom "map_mpoly hom" by (rule comm_ring_hom_mpoly_map) + interpret mmp: map_poly_comm_monoid_add_hom "map_mpoly hom" .. + show ?thesis unfolding mpoly_to_mpoly_poly_def + apply (subst mmp.hom_Sum_any, force) + apply (rule Sum_any.cong) + apply (unfold mp.map_poly_hom_monom map_mpoly_hom_monom) + by auto +qed +end + +context inj_comm_ring_hom +begin +lemma inj_comm_ring_hom_mpoly_map: "inj_comm_ring_hom (map_mpoly hom)" +proof - + interpret mp: comm_ring_hom "map_mpoly hom" by (rule comm_ring_hom_mpoly_map) + show ?thesis + proof (unfold_locales) + fix x + assume 0: "map_mpoly hom x = 0" + show "x = 0" + proof (intro mpoly_eqI) + fix m + show "MPoly_Type.coeff x m = MPoly_Type.coeff 0 m" + using arg_cong[OF 0, of "\ p. MPoly_Type.coeff p m"] by simp + qed + qed +qed + +lemma resultant_mpoly_poly_hom: "resultant_mpoly_poly x (map_mpoly hom p) (map_poly hom q) = map_mpoly hom (resultant_mpoly_poly x p q)" +proof - + interpret mp: inj_comm_ring_hom "map_mpoly hom" by (rule inj_comm_ring_hom_mpoly_map) + show ?thesis + unfolding resultant_mpoly_poly_def + unfolding mpoly_to_mpoly_poly_map_mpoly_hom + apply (subst mp.resultant_map_poly[symmetric]) + subgoal by (subst mp.degree_map_poly_hom, unfold_locales, auto) + subgoal by (subst mp.degree_map_poly_hom, unfold_locales, auto) + subgoal + apply (rule arg_cong[of _ _ "resultant _"], intro poly_eqI) + apply (subst coeff_map_poly, force)+ + by (simp add: map_mpoly_hom_Const) + done +qed +end + +lemma map_insort_key: assumes [simp]: "\ x y. g1 x \ g1 y \ g2 (f x) \ g2 (f y)" + shows "map f (insort_key g1 a xs) = insort_key g2 (f a) (map f xs)" + by (induct xs, auto) + +lemma map_sort_key: assumes [simp]: "\ x y. g1 x \ g1 y \ g2 (f x) \ g2 (f y)" + shows "map f (sort_key g1 xs) = sort_key g2 (map f xs)" + by (induct xs, auto simp: map_insort_key) + +hide_const (open) MPoly_Type.degree +hide_const (open) MPoly_Type.coeffs +hide_const (open) MPoly_Type.coeff +hide_const (open) Symmetric_Polynomials.lead_coeff + +subsection \More Facts about Resultants\ + +lemma resultant_iff_coprime_main: + fixes f g :: "'a :: field poly" + assumes deg: "degree f > 0 \ degree g > 0" +shows "resultant f g = 0 \ \ coprime f g" +proof (cases "resultant f g = 0") + case True + from resultant_zero_imp_common_factor[OF deg True] True + show ?thesis by simp +next + case False + from deg have fg: "f \ 0 \ g \ 0" by auto + from resultant_non_zero_imp_coprime[OF False fg] deg False + show ?thesis by auto +qed + +lemma resultant_zero_iff_coprime: fixes f g :: "'a :: field poly" + assumes "f \ 0 \ g \ 0" + shows "resultant f g = 0 \ \ coprime f g" +proof (cases "degree f > 0 \ degree g > 0") + case True + thus ?thesis using resultant_iff_coprime_main[OF True] by simp +next + case False + hence "degree f = 0" "degree g = 0" by auto + then obtain c d where f: "f = [:c:]" and g: "g = [:d:]" using degree0_coeffs by metis+ + from assms have cd: "c \ 0 \ d \ 0" unfolding f g by auto + have res: "resultant f g = 1" unfolding f g resultant_const by auto + have "coprime f g" + by (metis assms one_neq_zero res resultant_non_zero_imp_coprime) + with res show ?thesis by auto +qed + +text \The problem with the upcoming lemma is that "root" and "irreducibility" refer to the same type. + In the actual application we interested in "irreducibility" over the integers, but the roots + we are interested in are either real or complex.\ +lemma resultant_zero_iff_common_root_irreducible: fixes f g :: "'a :: field poly" + assumes irr: "irreducible g" + and root: "poly g a = 0" (* g has at least some root *) +shows "resultant f g = 0 \ (\ x. poly f x = 0 \ poly g x = 0)" +proof - + from irr root have deg: "degree g \ 0" using degree0_coeffs[of g] by fastforce + show ?thesis + proof + assume "\ x. poly f x = 0 \ poly g x = 0" + then obtain x where "poly f x = 0" "poly g x = 0" by auto + from resultant_zero[OF _ this] deg show "resultant f g = 0" by auto + next + assume "resultant f g = 0" + from resultant_zero_imp_common_factor[OF _ this] deg + have "\ coprime f g" by auto + from this[unfolded not_coprime_iff_common_factor] obtain r where + rf: "r dvd f" and rg: "r dvd g" and r: "\ is_unit r" by auto + from rg r irr have "g dvd r" + by (meson algebraic_semidom_class.irreducible_altdef) + with rf have "g dvd f" by auto + with root show "\ x. poly f x = 0 \ poly g x = 0" + by (intro exI[of _ a], auto simp: dvd_def) + qed +qed + + +lemma resultant_zero_iff_common_root_complex: fixes f g :: "complex poly" + assumes g: "g \ 0" +shows "resultant f g = 0 \ (\ x. poly f x = 0 \ poly g x = 0)" +proof (cases "degree g = 0") + case deg: False + show ?thesis + proof + assume "\ x. poly f x = 0 \ poly g x = 0" + then obtain x where "poly f x = 0" "poly g x = 0" by auto + from resultant_zero[OF _ this] deg show "resultant f g = 0" by auto + next + assume "resultant f g = 0" + from resultant_zero_imp_common_factor[OF _ this] deg + have "\ coprime f g" by auto + from this[unfolded not_coprime_iff_common_factor] obtain r where + rf: "r dvd f" and rg: "r dvd g" and r: "\ is_unit r" by auto + from rg g have r0: "r \ 0" by auto + with r have degr: "degree r \ 0" by simp + hence "\ constant (poly r)" + by (simp add: constant_degree) + from fundamental_theorem_of_algebra[OF this] obtain a where root: "poly r a = 0" by auto + from rf rg root show "\ x. poly f x = 0 \ poly g x = 0" + by (intro exI[of _ a], auto simp: dvd_def) + qed +next + case deg: True + from degree0_coeffs[OF deg] obtain c where gc: "g = [:c:]" by auto + from gc g have c: "c \ 0" by auto + hence "resultant f g \ 0" unfolding gc resultant_const by simp + with gc c show ?thesis by auto +qed + +subsection \Systems of Polynomials\ + +text \Definition of solving a system of polynomials, one being multivariate\ +definition mpoly_polys_solution :: "'a :: field mpoly \ (nat \ 'a poly) \ nat set \ (nat \ 'a) \ bool" where + "mpoly_polys_solution p qs N \ = ( + insertion \ p = 0 \ + (\ i \ N. poly (qs i) (\ (Suc i)) = 0))" + +text \The upcoming lemma shows how to eliminate single variables in multi-variate root-problems. + Because of the problem mentioned in @{thm [source] resultant_zero_iff_common_root_irreducible}, + we here restrict to polynomials over the complex numbers. Since the result computations are homomorphisms, + we are able to lift it to integer polynomials where we are interested in real or complex + roots.\ +lemma resultant_mpoly_polys_solution: fixes p :: "complex mpoly" + assumes nz: "0 \ qs ` N" + and i: "i \ N" +shows "mpoly_polys_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) \ + \ (\ v. mpoly_polys_solution p qs N (\((Suc i) := v)))" +proof - + let ?x = "Suc i" + let ?q = "qs i" + let ?mres = "resultant_mpoly_poly ?x p ?q" + from i obtain M where N: "N = insert i M" and MN: "M = N - {i}" and iM: "i \ M" by auto + from nz i have nzq: "?q \ 0" by auto + hence lc0: "lead_coeff (qs i) \ 0" by auto + have "mpoly_polys_solution ?mres qs (N - {i}) \ \ + insertion \ ?mres = 0 \ (\ i \ M. poly (qs i) (\ (Suc i)) = 0)" + unfolding mpoly_polys_solution_def MN .. + also have "insertion \ ?mres = 0 \ resultant (partial_insertion \ ?x p) ?q = 0" + by (rule insertion_resultant_mpoly_poly_zero[OF nzq]) + also have "\ \ (\v. poly (partial_insertion \ ?x p) v = 0 \ poly ?q v = 0)" + by (rule resultant_zero_iff_common_root_complex[OF nzq]) + also have "\ \ (\v. insertion (\(?x := v)) p = 0 \ poly ?q v = 0)" (is "?lhs = ?rhs") + proof (intro iff_exI conj_cong refl arg_cong[of _ _ "\ x. x = 0"]) + fix v + have "poly (partial_insertion \ ?x p) v = poly (partial_insertion \ ?x p) ((\(?x := v)) ?x)" by simp + also have "\ = insertion (\(?x := v)) p" + by (rule insertion_partial_insertion, auto) + finally show "poly (partial_insertion \ ?x p) v = insertion (\(?x := v)) p" . + qed + also have "\ \ (\i\M. poly (qs i) (\ (Suc i)) = 0) + \ (\v. insertion (\(?x := v)) p = 0 \ poly (qs i) v = 0 \ (\i\M. poly (qs i) ((\(?x := v)) (Suc i)) = 0))" + using iM by auto + also have "\ \ (\ v. mpoly_polys_solution p qs N (\((Suc i) := v)))" + unfolding mpoly_polys_solution_def N by (intro iff_exI, auto) + finally + show ?thesis . +qed + +text \We now restrict solutions to be evaluated to zero outside the variable range. Then there are only finitely + many solutions for our applications.\ +definition mpoly_polys_zero_solution :: "'a :: field mpoly \ (nat \ 'a poly) \ nat set \ (nat \ 'a) \ bool" where + "mpoly_polys_zero_solution p qs N \ = (mpoly_polys_solution p qs N \ + \ (\ i. i \ insert 0 (Suc ` N) \ \ i = 0))" + +lemma resultant_mpoly_polys_zero_solution: fixes p :: "complex mpoly" + assumes nz: "0 \ qs ` N" + and i: "i \ N" +shows + "mpoly_polys_zero_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) \ + \ \ v. mpoly_polys_zero_solution p qs N (\(Suc i := v))" + "mpoly_polys_zero_solution p qs N \ + \ mpoly_polys_zero_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) (\(Suc i := 0))" +proof - + assume "mpoly_polys_zero_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) \" + hence 1: "mpoly_polys_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) \" and 2: "(\ i. i \ insert 0 (Suc ` (N - {i})) \ \ i = 0)" + unfolding mpoly_polys_zero_solution_def by auto + from resultant_mpoly_polys_solution[of qs N _ p \, OF nz i] 1 obtain v where "mpoly_polys_solution p qs N (\(Suc i := v))" by auto + with 2 have "mpoly_polys_zero_solution p qs N (\(Suc i := v))" using i unfolding mpoly_polys_zero_solution_def by auto + thus "\ v. mpoly_polys_zero_solution p qs N (\(Suc i := v))" .. +next + assume "mpoly_polys_zero_solution p qs N \" + from this[unfolded mpoly_polys_zero_solution_def] have 1: "mpoly_polys_solution p qs N \" and 2: "\i. i \ insert 0 (Suc ` N) \ \ i = 0" by auto + from 1 have "mpoly_polys_solution p qs N (\(Suc i := \ (Suc i)))" by auto + hence "\ v. mpoly_polys_solution p qs N (\(Suc i := v))" by blast + with resultant_mpoly_polys_solution[of qs N _ p \, OF nz i] have "mpoly_polys_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) \" by auto + hence "mpoly_polys_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) (\ (Suc i := 0))" + unfolding mpoly_polys_solution_def + apply simp + apply (subst insertion_irrelevant_vars[of _ _ \]) + by (insert vars_resultant_mpoly_poly, auto) + thus "mpoly_polys_zero_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) (\(Suc i := 0))" + unfolding mpoly_polys_zero_solution_def using 2 by auto +qed + +text \The following two lemmas show that if we start with a system of polynomials with finitely + many solutions, then the resulting polynomial cannot be the zero-polynomial.\ +lemma finite_resultant_mpoly_polys_non_empty: fixes p :: "complex mpoly" + assumes nz: "0 \ qs ` N" + and i: "i \ N" + and fin: "finite {\. mpoly_polys_zero_solution p qs N \}" +shows "finite {\. mpoly_polys_zero_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i}) \}" +proof - + let ?solN = "mpoly_polys_zero_solution p qs N" + let ?solN1 = "mpoly_polys_zero_solution (resultant_mpoly_poly (Suc i) p (qs i)) qs (N - {i})" + let ?x = "Suc i" + note defs = mpoly_polys_zero_solution_def + define zero where "zero \ = \(?x := 0)" for \ :: "nat \ complex" + { + fix \ + assume sol: "?solN1 \" + from sol[unfolded defs] have 0: "\ ?x = 0" by auto + from resultant_mpoly_polys_zero_solution(1)[of qs N i p, OF nz i sol] obtain v + where "?solN (\(?x := v))" by auto + hence sol: "\(?x := v) \ {\. ?solN \}" by auto + hence "zero (\(?x := v)) \ zero ` {\. ?solN \}" by auto + also have "zero (\(?x := v)) = \" using 0 by (auto simp: zero_def) + finally have "\ \ zero ` {\. ?solN \}" . + } + hence "{\. ?solN1 \} \ zero ` {\. ?solN \}" by blast + from finite_subset[OF this finite_imageI[OF fin]] + show ?thesis . +qed + +lemma finite_resultant_mpoly_polys_empty: fixes p :: "complex mpoly" + assumes "finite {\. mpoly_polys_zero_solution p qs {} \}" + shows "p \ 0" +proof + define g where "g x = (\ i :: nat. if i = 0 then x else 0)" for x :: complex + assume "p = 0" + hence "\ x. mpoly_polys_zero_solution p qs {} (g x)" + unfolding mpoly_polys_zero_solution_def mpoly_polys_solution_def g_def by auto + hence "range g \ {\. mpoly_polys_zero_solution p qs {} \}" by auto + from finite_subset[OF this assms] have "finite (range g)" . + moreover have "inj g" unfolding g_def inj_on_def by metis + ultimately have "finite (UNIV :: complex set)" by simp + thus False using infinite_UNIV_char_0 by auto +qed + +subsection \Elimination of Auxiliary Variables\ + +fun eliminate_aux_vars :: "'a :: comm_ring_1 mpoly \ (nat \ 'a poly) \ nat list \ 'a poly" where + "eliminate_aux_vars p qs [] = mpoly_to_poly 0 p" +| "eliminate_aux_vars p qs (i # is) = eliminate_aux_vars (resultant_mpoly_poly (Suc i) p (qs i)) qs is" + + +lemma eliminate_aux_vars_of_int_poly: + "eliminate_aux_vars (map_mpoly (of_int :: _ \ 'a :: {comm_ring_1,ring_char_0}) mp) (of_int_poly \ qs) is + = of_int_poly (eliminate_aux_vars mp qs is)" +proof - + let ?h = "of_int :: _ \ 'a" + interpret mp: comm_ring_hom "(map_mpoly ?h)" + by (rule of_int_hom.comm_ring_hom_mpoly_map) + show ?thesis + proof (induct "is" arbitrary: mp) + case Nil + show ?case by (simp add: of_int_hom.mpoly_to_poly_map_mpoly_hom) + next + case (Cons i "is" mp) + show ?case unfolding eliminate_aux_vars.simps Cons[symmetric] + apply (rule arg_cong[of _ _ "\ x. eliminate_aux_vars x _ _"], unfold o_def) + by (rule of_int_hom.resultant_mpoly_poly_hom) + qed +qed + +text \The polynomial of the elimination process will represent the first value @{term "\ 0 :: complex"} of any + solution to the multi-polynomial problem.\ +lemma eliminate_aux_vars: fixes p :: "complex mpoly" + assumes "distinct is" + and "vars p \ insert 0 (Suc ` set is)" + and "finite {\. mpoly_polys_zero_solution p qs (set is) \}" + and "0 \ qs ` set is" + and "mpoly_polys_solution p qs (set is) \" +shows "poly (eliminate_aux_vars p qs is) (\ 0) = 0 \ eliminate_aux_vars p qs is \ 0" + using assms +proof (induct "is" arbitrary: p) + case (Nil p) + from Nil(3) finite_resultant_mpoly_polys_empty[of p] + have p0: "p \ 0" by auto + from Nil(2) have vars: "vars p \ {0}" by auto + note [simp] = poly_eq_insertion[OF this] + from Nil(5)[unfolded mpoly_polys_solution_def] + have "insertion \ p = 0" by auto + also have "insertion \ p = insertion (\v. \ 0) p" + by (rule insertion_irrelevant_vars, insert vars, auto) + finally + show ?case using p0 mpoly_to_poly_inverse[OF vars] by (auto simp: poly_to_mpoly0) +next + case (Cons i "is" p) + let ?x = "Suc i" + let ?p = "resultant_mpoly_poly ?x p (qs i)" + have dist: "distinct is" using Cons(2) by auto + have vars: "vars ?p \ insert 0 (Suc ` set is)" using Cons(3) vars_resultant_mpoly_poly[of ?x p "qs i"] by auto + have fin: "finite {\. mpoly_polys_zero_solution ?p qs (set is) \}" + using finite_resultant_mpoly_polys_non_empty[of qs "set (i # is)" i p, OF Cons(5)] Cons(2,4) by auto + have 0: "0 \ qs ` set is" using Cons(5) by auto + have "(\v. mpoly_polys_solution p qs (set (i # is)) (\(?x := v)))" + using Cons(6) by (intro exI[of _ "\ ?x"], auto) + from this resultant_mpoly_polys_solution[OF Cons(5), of i p \] + have "mpoly_polys_solution ?p qs (set (i # is) - {i}) \" + by auto + also have "set (i # is) - {i} = set is" using Cons(2) by auto + finally have "mpoly_polys_solution ?p qs (set is) \" by auto + note IH = Cons(1)[OF dist vars fin 0 this] + show ?case unfolding eliminate_aux_vars.simps using IH by simp +qed + +subsection \A Representing Polynomial for the Roots of a Polynomial with Algebraic Coefficients\ + +text \First convert an algebraic polynomial into a system of integer polynomials.\ +definition initial_root_problem :: "'a :: {is_rat,field_gcd} poly \ int mpoly \ (nat \ 'a \ int poly) list" where + "initial_root_problem p = (let + n = degree p; + cs = coeffs p; + rcs = remdups (filter (\ c. c \ \) cs); + pairs = map (\ c. (c, min_int_poly c)) rcs; + spairs = sort_key (\ (c,f). degree f) pairs; \ \sort by degree so that easy computations will be done first\ + triples = zip [0 ..< length spairs] spairs; + mpoly = (sum (\ i. let c = coeff p i in + MPoly_Type.monom (Poly_Mapping.single 0 i) 1 * \ \$x_0 ^ i * ...$\ + (case find (\ (j,d,f). d = c) triples of + None \ Const (to_int c) + | Some (j,pair) \ Var (Suc j))) + {..n}) + in (mpoly, triples))" + +text \And then eliminate all auxiliary variables\ + +definition representative_poly :: "'a :: {is_rat,field_char_0,field_gcd} poly \ int poly" where + "representative_poly p = (case initial_root_problem p of + (mp, triples) \ + let is = map fst triples; + qs = (\ j. snd (snd (triples ! j))) + in eliminate_aux_vars mp qs is)" + + +subsection \Soundness Proof for Complex Algebraic Polynomials\ + +lemma get_representative_complex: fixes p :: "complex poly" + assumes p: "p \ 0" + and algebraic: "Ball (set (coeffs p)) algebraic" + and res: "initial_root_problem p = (mp, triples)" + and "is": "is = map fst triples" + and qs: "\ j. j < length is \ qs j = snd (snd (triples ! j))" + and root: "poly p x = 0" +shows "eliminate_aux_vars mp qs is represents x" +proof - + define rcs where "rcs = remdups (filter (\c. c \ \) (coeffs p))" + define spairs where "spairs = sort_key (\(c, f). degree f) (map (\c. (c, min_int_poly c)) rcs)" + let ?find = "\ i. find (\(j, d, f). d = coeff p i) triples" + define trans where "trans i = (case ?find i of None \ Const (to_int (coeff p i)) + | Some (j, pair) \ Var (Suc j))" for i + note res = res[unfolded initial_root_problem_def Let_def, folded rcs_def, folded spairs_def] + have triples: "triples = zip [0..i\degree p. MPoly_Type.monom (Poly_Mapping.single 0 i) 1 * trans i)" using res by auto + have dist_rcs: "distinct rcs" unfolding rcs_def by auto + hence "distinct (map fst (map (\c. (c, min_int_poly c)) rcs))" by (simp add: o_def) + hence dist_spairs: "distinct (map fst spairs)" unfolding spairs_def + by (metis (no_types, lifting) distinct_map distinct_sort set_sort) + { + fix c + assume "c \ set rcs" + hence "c \ set (coeffs p)" unfolding rcs_def by auto + with algebraic have "algebraic c" by auto + } note rcs_alg = this + { + fix c + assume c: "c \ range (coeff p)" "c \ \" + hence "c \ set (coeffs p)" unfolding range_coeff by auto + with c have crcs: "c \ set rcs" unfolding rcs_def by auto + from rcs_alg[OF crcs] have "algebraic c" . + from min_int_poly_represents[OF this] + have "min_int_poly c represents c" . + hence "\ f. (c,f) \ set spairs \ f represents c" using crcs unfolding spairs_def by auto + } + have dist_is: "distinct is" unfolding "is" triples by simp + note eliminate = eliminate_aux_vars[OF dist_is] + let ?mp = "map_mpoly of_int mp :: complex mpoly" + have vars_mp: "vars mp \ insert 0 (Suc ` set is)" + unfolding mp + apply (rule order.trans[OF vars_setsum], force) + apply (rule UN_least, rule order.trans[OF vars_mult], rule Un_least) + apply (intro order.trans[OF vars_monom_single], force) + subgoal for i + proof - + show ?thesis + proof (cases "?find i") + case None + show ?thesis unfolding trans_def None by auto + next + case (Some j_pair) + then obtain j c f where find: "?find i = Some (j,c,f)" by (cases j_pair, auto) + from find_Some_D[OF find] have "Suc j \ Suc ` (fst ` set triples)" by force + thus ?thesis unfolding trans_def find by (simp add: vars_Var "is") + qed + qed + done + hence varsMp: "vars ?mp \ insert 0 (Suc ` set is)" using vars_map_mpoly_subset by auto + note eliminate = eliminate[OF this] + let ?f = "\ j. snd (snd (triples ! j))" + let ?c = "\ j. fst (snd (triples ! j))" + { + fix j + assume "j \ set is" + hence "(?c j, ?f j) \ set spairs" unfolding "is" triples by simp + hence "?f j represents ?c j" "?f j = min_int_poly (?c j)" unfolding spairs_def + by (auto intro: min_int_poly_represents[OF rcs_alg]) + } note is_repr = this + let ?qs = "(of_int_poly o qs) :: nat \ complex poly" + { + fix j + assume "j \ set is" + hence "j < length is" unfolding "is" triples by simp + } note j_len = this + have qs_0: "0 \ qs ` set is" + proof + assume "0 \ qs ` set is" + then obtain j where j: "j \ set is" and 0: "qs j = 0" by auto + from is_repr[OF j] have "?f j \ 0" by auto + with 0 show False unfolding qs[OF j_len[OF j]] by auto + qed + hence qs0: "0 \ ?qs ` set is" by auto + note eliminate = eliminate[OF _ this] + define roots where "roots p = (SOME xs. set xs = {x . poly p x = 0})" for p :: "complex poly" + { + fix p :: "complex poly" + assume "p \ 0" + from someI_ex[OF finite_list[OF poly_roots_finite[OF this]], folded roots_def] + have "set (roots p) = {x. poly p x = 0}" . + } note roots = this + define qs_roots where "qs_roots = concat_lists (map (\ i. roots (?qs i)) [0 ..< length triples])" + define evals where "evals = concat (map (\ part. let + q = partial_insertion (\ i. part ! (i - 1)) 0 ?mp; + new_roots = roots q + in map (\ r. r # part) new_roots) qs_roots)" + define conv where "conv roots i = (if i \ length triples then roots ! i else 0 :: complex)" for roots i + define alphas where "alphas = map conv evals" + { + fix n + assume n: "n \ {..degree p}" + let ?cn = "coeff p n" + from n have mem: "?cn \ set (coeffs p)" using p unfolding Polynomial.coeffs_def by force + { + assume "?cn \ \" + with mem have "?cn \ set rcs" unfolding rcs_def by auto + hence "(?cn, min_int_poly ?cn) \ set spairs" unfolding spairs_def by auto + hence "\ i. (i, ?cn, min_int_poly ?cn) \ set triples" unfolding triples set_zip set_conv_nth + by force + hence "?find n \ None" unfolding find_None_iff by auto + } + } note non_int_find = this + have fin: "finite {\. mpoly_polys_zero_solution ?mp ?qs (set is) \}" + proof (rule finite_subset[OF _ finite_set[of alphas]], standard, clarify) + fix \ + assume sol: "mpoly_polys_zero_solution ?mp ?qs (set is) \" + define part where "part = map (\ i. \ (Suc i)) [0 ..< length triples]" + { + fix i + assume "i > length triples" + hence "i \ insert 0 (Suc ` set is)" unfolding triples "is" by auto + hence "\ i = 0" using sol[unfolded mpoly_polys_zero_solution_def] by auto + } note alpha0 = this + { + fix i + assume "i < length triples" + hence i: "i \ set is" unfolding triples "is" by auto + from qs0 i have 0: "?qs i \ 0" by auto + from i sol[unfolded mpoly_polys_zero_solution_def mpoly_polys_solution_def] + have "poly (?qs i) (\ (Suc i)) = 0" by auto + hence "\ (Suc i) \ set (roots (?qs i))" "poly (?qs i) (\ (Suc i)) = 0" using roots[OF 0] by auto + } note roots2 = this + hence part: "part \ set qs_roots" + unfolding part_def qs_roots_def concat_lists_listset listset by auto + let ?gamma = "(\i. part ! (i - 1))" + let ?f = "partial_insertion ?gamma 0 ?mp" + have "\ 0 \ set (roots ?f)" + proof - + from sol[unfolded mpoly_polys_zero_solution_def mpoly_polys_solution_def] + have "0 = insertion \ ?mp" by simp + also have "\ = insertion (\ i. if i \ length triples then \ i else part ! (i - 1)) ?mp" + (is "_ = insertion ?beta _") + proof (rule insertion_irrelevant_vars) + fix i + assume "i \ vars ?mp" + from set_mp[OF varsMp this] have "i \ length triples" unfolding triples "is" by auto + thus "\ i = ?beta i" by auto + qed + also have "\ = poly (partial_insertion (?beta(0 := part ! 0)) 0 ?mp) (?beta 0)" + by (subst insertion_partial_insertion, auto) + also have "?beta(0 := part ! 0) = ?gamma" unfolding part_def + by (intro ext, auto) + finally have root: "poly ?f (\ 0) = 0" by auto + have "?f \ 0" + proof + interpret mp: inj_comm_ring_hom "map_mpoly complex_of_int" + by (rule of_int_hom.inj_comm_ring_hom_mpoly_map) + assume "?f = 0" + hence "0 = coeff ?f (degree p)" by simp + also have "\ = insertion ?gamma (coeff (mpoly_to_mpoly_poly 0 ?mp) (degree p))" + unfolding insertion_coeff_mpoly_to_mpoly_poly[symmetric] .. + also have "coeff (mpoly_to_mpoly_poly 0 ?mp) (degree p) = map_mpoly of_int (coeff (mpoly_to_mpoly_poly 0 mp) (degree p))" + unfolding of_int_hom.mpoly_to_mpoly_poly_map_mpoly_hom + by (subst coeff_map_poly, auto) + also have "coeff (mpoly_to_mpoly_poly 0 mp) (degree p) = + (\x. MPoly_Type.monom (remove_key 0 x) (MPoly_Type.coeff mp x) when lookup x 0 = degree p)" + unfolding mpoly_to_mpoly_poly_def when_def + by (subst coeff_hom.hom_Sum_any, force, unfold Polynomial.coeff_monom, auto) + also have "\ = (\x. MPoly_Type.monom (remove_key 0 x) + (\xa\degree p. let xx = Poly_Mapping.single 0 xa in + \(a, b). MPoly_Type.coeff (trans xa) b when x = xx + b when + a = xx) when + lookup x 0 = degree p)" unfolding mp coeff_sum More_MPoly_Type.coeff_monom coeff_mpoly_times Let_def + apply (subst prod_fun_unfold_prod, force, force) + apply (unfold when_mult, subst when_commute) + by (auto simp: when_def intro!: Sum_any.cong sum.cong if_cong arg_cong[of _ _ "MPoly_Type.monom _"]) + also have "\ = (\x. MPoly_Type.monom (remove_key 0 x) + (\i\degree p. \m. MPoly_Type.coeff (trans i) m when x = Poly_Mapping.single 0 i + m) when + lookup x 0 = degree p)" + unfolding Sum_any_when_dependent_prod_left Let_def by simp + also have "\ = (\x. MPoly_Type.monom (remove_key 0 x) + (\i \ {degree p}. \m. MPoly_Type.coeff (trans i) m when x = Poly_Mapping.single 0 i + m) when + lookup x 0 = degree p)" + apply (intro Sum_any.cong when_cong refl arg_cong[of _ _ "MPoly_Type.monom _"] sum.mono_neutral_right, force+) + apply (intro ballI Sum_any_zeroI, auto simp: when_def) + subgoal for i x + proof (goal_cases) + case 1 + hence "lookup x 0 > 0" by (auto simp: lookup_add) + moreover have "0 \ vars (trans i)" unfolding trans_def + by (auto split: option.splits simp: vars_Var) + ultimately show ?thesis + by (metis set_mp coeff_notin_vars in_keys_iff neq0_conv) + qed + done + also have "\ = (\x. MPoly_Type.monom (remove_key 0 x) + (\m. MPoly_Type.coeff (trans (degree p)) m when x = Poly_Mapping.single 0 (degree p) + m) when + lookup x 0 = degree p)" (is "_ = ?mid") + by simp + also have "insertion ?gamma (map_mpoly of_int \) \ 0" + proof (cases "?find (degree p)") + case None + from non_int_find[of "degree p"] None + have lcZ: "lead_coeff p \ \" by auto + have "?mid = (\x. MPoly_Type.monom (remove_key 0 x) + (\m. (to_int (lead_coeff p) when + x = Poly_Mapping.single 0 (degree p) + m when m = 0)) when + lookup x 0 = degree p)" + using None unfolding trans_def None option.simps mpoly_coeff_Const when_def + by (intro Sum_any.cong if_cong refl, intro arg_cong[of _ _ "MPoly_Type.monom _"] Sum_any.cong, auto) + also have "\ = (\x. MPoly_Type.monom (remove_key 0 x) + (to_int (lead_coeff p) when x = Poly_Mapping.single 0 (degree p)) when + lookup x 0 = degree p when x = Poly_Mapping.single 0 (degree p))" + unfolding Sum_any_when_equal[of _ 0] + by (intro Sum_any.cong, auto simp: when_def) + also have "\ = MPoly_Type.monom (remove_key 0 (Poly_Mapping.single 0 (degree p))) + (to_int (lead_coeff p)) " + unfolding Sum_any_when_equal by simp + also have "\ = Const (to_int (lead_coeff p))" by (simp add: mpoly_monom_0_eq_Const) + also have "map_mpoly of_int \ = Const (lead_coeff p)" + unfolding of_int_hom.map_mpoly_hom_Const of_int_to_int[OF lcZ] by simp + also have "insertion ?gamma \ = lead_coeff p" by simp + also have "\ \ 0" using p by auto + finally show ?thesis . + next + case Some + from find_Some_D[OF this] Some obtain j f where mem: "(j,lead_coeff p,f) \ set triples" and + Some: "?find (degree p) = Some (j, lead_coeff p, f)" by auto + from mem have j: "j < length triples" unfolding triples set_zip by auto + have "?mid = (\x. if lookup x 0 = degree p + then MPoly_Type.monom (remove_key 0 x) + (\m. 1 when m = Poly_Mapping.single (Suc j) 1 when x = Poly_Mapping.single 0 (degree p) + m) + else 0)" + unfolding trans_def Some option.simps split when_def coeff_Var by auto + also have "\ = (\x. if lookup x 0 = degree p + then MPoly_Type.monom (remove_key 0 x) 1 + when x = Poly_Mapping.single 0 (degree p) + Poly_Mapping.single (Suc j) 1 + else 0 when x = Poly_Mapping.single 0 (degree p) + Poly_Mapping.single (Suc j) 1)" + apply (subst when_commute) + apply (unfold Sum_any_when_equal) + by (rule Sum_any.cong, auto simp: when_def) + also have "\ = (\x. (MPoly_Type.monom (remove_key 0 x) 1 when lookup x 0 = degree p) + when x = Poly_Mapping.single 0 (degree p) + Poly_Mapping.single (Suc j) 1)" + by (rule Sum_any.cong, auto simp: when_def) + also have "\ = MPoly_Type.monom (Poly_Mapping.single (Suc j) 1) 1" + unfolding Sum_any_when_equal unfolding when_def + by (simp add: lookup_add remove_key_add[symmetric] + remove_key_single' lookup_single) + also have "\ = Var (Suc j)" + by (intro mpoly_eqI, simp add: coeff_Var coeff_monom) + also have "map_mpoly complex_of_int \ = Var (Suc j)" + by (simp add: map_mpoly_Var) + also have "insertion ?gamma \ = part ! j" by simp + also have "\ = \ (Suc j)" unfolding part_def using j by auto + also have "\ \ 0" + proof + assume "\ (Suc j) = 0" + with roots2(2)[OF j] have root0: "poly (?qs j) 0 = 0" by auto + from j "is" have ji: "j < length is" by auto + hence jis: "j \ set is" unfolding "is" triples set_zip by auto + from mem have tj: "triples ! j = (j, lead_coeff p, f)" unfolding triples set_zip by auto + from root0[unfolded qs[OF ji] o_def tj] + have rootf: "poly f 0 = 0" by auto + from is_repr[OF jis, unfolded tj] have rootlc: "ipoly f (lead_coeff p) = 0" + and f: "f = min_int_poly (lead_coeff p)" by auto + from f have irr: "irreducible f" by auto + from rootf have "[:0,1:] dvd f" using dvd_iff_poly_eq_0 by fastforce + from this[unfolded dvd_def] obtain g where f: "f = [:0, 1:] * g" by auto + from irreducibleD[OF irr f] have "is_unit g" + by (metis is_unit_poly_iff one_neq_zero one_pCons pCons_eq_iff) + then obtain c where g: "g = [:c:]" and c: "c dvd 1" unfolding is_unit_poly_iff by auto + from rootlc[unfolded f g] c have "lead_coeff p = 0" by auto + with p show False by auto + qed + finally show ?thesis . + qed + finally show False by auto + qed + from roots[OF this] root show ?thesis by auto + qed + hence "\ 0 # part \ set evals" + unfolding evals_def set_concat Let_def set_map + by (auto intro!: bexI[OF _ part]) + hence "map \ [0 ..< Suc (length triples)] \ set evals" unfolding part_def + by (metis Utility.map_upt_Suc) + hence "conv (map \ [0 ..< Suc (length triples)]) \ set alphas" unfolding alphas_def by auto + also have "conv (map \ [0 ..< Suc (length triples)]) = \" + proof + fix i + show "conv (map \ [0.. i" + unfolding conv_def using alpha0 + by (cases "i < length triples"; cases "i = length triples"; auto simp: nth_append) + qed + finally show "\ \ set alphas" . + qed + note eliminate = eliminate[OF this] + define \ where "\ x j = (if j = 0 then x else ?c (j - 1))" for x j + have \: "\ x (Suc j) = ?c j" "\ x 0 = x" for j x unfolding \_def by auto + interpret mp: inj_comm_ring_hom "map_mpoly complex_of_int" by (rule of_int_hom.inj_comm_ring_hom_mpoly_map) + have ins: "insertion (\ x) ?mp = poly p x" for x + unfolding poly_altdef mp mp.hom_sum insertion_sum insertion_mult mp.hom_mult + proof (rule sum.cong[OF refl], subst mult.commute, rule arg_cong2[of _ _ _ _ "(*)"]) + fix n + assume n: "n \ {..degree p}" + let ?cn = "coeff p n" + from n have mem: "?cn \ set (coeffs p)" using p unfolding Polynomial.coeffs_def by force + have "insertion (\ x) (map_mpoly complex_of_int (MPoly_Type.monom (Poly_Mapping.single 0 n) 1)) = (\a. \ x a ^ (n when a = 0))" + unfolding of_int_hom.map_mpoly_hom_monom by (simp add: lookup_single) + also have "\ = (\a. if a = 0 then \ x a ^ n else 1)" + by (rule Prod_any.cong, auto simp: when_def) + also have "\ = \ x 0 ^ n" by simp + also have "\ = x ^ n" unfolding \ .. + finally show "insertion (\ x) (map_mpoly complex_of_int (MPoly_Type.monom (Poly_Mapping.single 0 n) 1)) = x ^ n" . + show "insertion (\ x) (map_mpoly complex_of_int (trans n)) = ?cn" + proof (cases "?find n") + case None + with non_int_find[OF n] have ints: "?cn \ \" by auto + from None show ?thesis unfolding trans_def using ints + by (simp add: of_int_hom.map_mpoly_hom_Const of_int_to_int) + next + case (Some triple) + from find_Some_D[OF this] this obtain j f + where mem: "(j,?cn,f) \ set triples" and Some: "?find n = Some (j,?cn,f)" + by (cases triple, auto) + from mem have "triples ! j = (j,?cn,f)" unfolding triples set_zip by auto + thus ?thesis unfolding trans_def Some by (simp add: map_mpoly_Var \_def) + qed + qed + from root have "insertion (\ x) ?mp = 0" unfolding ins by auto + hence "mpoly_polys_solution ?mp ?qs (set is) (\ x)" + unfolding mpoly_polys_solution_def + proof (standard, intro ballI) + fix j + assume j: "j \ set is" + from is_repr[OF this] + show "poly (?qs j) (\ x (Suc j)) = 0" unfolding \ qs[OF j_len[OF j]] o_def by auto + qed + note eliminate = eliminate[OF this, unfolded \ eliminate_aux_vars_of_int_poly] + thus "eliminate_aux_vars mp qs is represents x" by auto +qed + +lemma representative_poly_complex: fixes x :: complex + assumes p: "p \ 0" + and algebraic: "Ball (set (coeffs p)) algebraic" + and root: "poly p x = 0" + shows "representative_poly p represents x" +proof - + obtain mp triples where init: "initial_root_problem p = (mp, triples)" by force + from get_representative_complex[OF p algebraic init refl _ root] + show ?thesis unfolding representative_poly_def init Let_def by auto +qed + +subsection \Soundness Proof for Real Algebraic Polynomials\ + +text \We basically use the result for complex algebraic polynomials which + are a superset of real algebraic polynomials.\ + + +lemma initial_root_problem_complex_of_real_poly: + "initial_root_problem (map_poly complex_of_real p) = + map_prod id (map (map_prod id (map_prod complex_of_real id))) (initial_root_problem p)" +proof - + let ?c = "of_real :: real \ complex" + let ?cp = "map_poly ?c" + let ?p = "?cp p :: complex poly" + define cn where "cn = degree ?p" + define n where "n = degree p" + have n: "cn = n" unfolding n_def cn_def by simp + note def = initial_root_problem_def[of ?p] + note def = def[folded cn_def, unfolded n] + define ccs where "ccs = coeffs ?p" + define cs where "cs = coeffs p" + have cs: "ccs = map ?c cs" + unfolding ccs_def cs_def by auto + note def = def[folded ccs_def] + define crcs where "crcs = remdups (filter (\c. c \ \) ccs)" + define rcs where "rcs = remdups (filter (\c. c \ \) cs)" + have rcs: "crcs = map ?c rcs" + unfolding crcs_def rcs_def cs by (induct cs, auto) + define cpairs where "cpairs = map (\c. (c, min_int_poly c)) crcs" + define pairs where "pairs = map (\c. (c, min_int_poly c)) rcs" + have pairs: "cpairs = map (map_prod ?c id) pairs" + unfolding pairs_def cpairs_def rcs by auto + define cspairs where "cspairs = sort_key (\(c, y). degree y) cpairs" + define spairs where "spairs = sort_key (\(c, y). degree y) pairs" + have spairs: "cspairs = map (map_prod ?c id) spairs" + unfolding spairs_def cspairs_def pairs + by (rule sym, rule map_sort_key, auto) + define ctriples where "ctriples = zip [0.. x. _ * x"], induct triples, auto) +qed + + +lemma representative_poly_real: fixes x :: real + assumes p: "p \ 0" + and algebraic: "Ball (set (coeffs p)) algebraic" + and root: "poly p x = 0" +shows "representative_poly p represents x" +proof - + obtain mp triples where init: "initial_root_problem p = (mp, triples)" by force + define "is" where "is = map fst triples" + define qs where "qs = (\ j. snd (snd (triples ! j)))" + let ?c = "of_real :: real \ complex" + let ?cp = "map_poly ?c" + let ?ct = "map (map_prod id (map_prod ?c id))" + let ?p = "?cp p :: complex poly" + have p: "?p \ 0" using p by auto + have "initial_root_problem ?p = map_prod id ?ct (initial_root_problem p)" + by (rule initial_root_problem_complex_of_real_poly) + from this[unfolded init] + have res: "initial_root_problem ?p = (mp, ?ct triples)" + by auto + from root have "0 = ?c (poly p x)" by simp + also have "\ = poly ?p (?c x)" by simp + finally have root: "poly ?p (?c x) = 0" by simp + have qs: "j < length is \ qs j = snd (snd (?ct triples ! j))" for j + unfolding is_def qs_def by (auto simp: set_conv_nth) + have "is": "is = map fst (?ct triples)" unfolding is_def by auto + { + fix cc + assume "cc \ set (coeffs ?p)" + then obtain c where "c \ set (coeffs p)" and cc: "cc = ?c c" by auto + from algebraic this(1) have "algebraic cc" + unfolding cc algebraic_complex_iff by auto + } + hence algebraic: "Ball (set (coeffs ?p)) algebraic" .. + from get_representative_complex[OF p this res "is" qs root] + have "eliminate_aux_vars mp qs is represents ?c x" . + hence "eliminate_aux_vars mp qs is represents x" by simp + thus ?thesis unfolding representative_poly_def res init split Let_def qs_def is_def . +qed + +subsection \Algebraic Closedness of Complex Algebraic Numbers\ + +(* TODO: could be generalised to arbitrary algebraically closed fields? *) +lemma complex_algebraic_numbers_are_algebraically_closed: + assumes nc: "\ constant (poly p)" + and alg: "Ball (set (coeffs p)) algebraic" + shows "\ z :: complex. algebraic z \ poly p z = 0" +proof - + from fundamental_theorem_of_algebra[OF nc] obtain z where + root: "poly p z = 0" by auto + from algebraic_representsI[OF representative_poly_complex[OF _ alg root]] nc root + have "algebraic z \ poly p z = 0" + using constant_degree degree_0 by blast + thus ?thesis .. +qed + + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Roots_of_Algebraic_Poly_Impl.thy b/thys/Factor_Algebraic_Polynomial/Roots_of_Algebraic_Poly_Impl.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Roots_of_Algebraic_Poly_Impl.thy @@ -0,0 +1,47 @@ +subsection \Executable Version to Compute Representative Polynomials\ + +theory Roots_of_Algebraic_Poly_Impl +imports + Roots_of_Algebraic_Poly + Polynomials.MPoly_Type_Class_FMap +begin + +text \We need to specialize our code to real and complex polynomials, + since @{const algebraic} and @{const min_int_poly} are + not executable in their parametric versions.\ + +definition initial_root_problem_real :: "real poly \ _" where + [simp]: "initial_root_problem_real p = initial_root_problem p" + +definition initial_root_problem_complex :: "complex poly \ _" where + [simp]: "initial_root_problem_complex p = initial_root_problem p" + +lemmas initial_root_problem_code = + initial_root_problem_real_def[unfolded initial_root_problem_def] + initial_root_problem_complex_def[unfolded initial_root_problem_def] + +declare initial_root_problem_code[code] + +lemma initial_root_problem_code_unfold[code_unfold]: + "initial_root_problem = initial_root_problem_complex" + "initial_root_problem = initial_root_problem_real" + by (intro ext, simp)+ + +definition representative_poly_real :: "real poly \ _" where + [simp]: "representative_poly_real p = representative_poly p" + +definition representative_poly_complex :: "complex poly \ _" where + [simp]: "representative_poly_complex p = representative_poly p" + +lemmas representative_poly_code = + representative_poly_real_def[unfolded representative_poly_def] + representative_poly_complex_def[unfolded representative_poly_def] + +declare representative_poly_code[code] + +lemma representative_poly_code_unfold[code_unfold]: + "representative_poly = representative_poly_complex" + "representative_poly = representative_poly_real" + by (intro ext, simp)+ + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Roots_of_Real_Complex_Poly.thy b/thys/Factor_Algebraic_Polynomial/Roots_of_Real_Complex_Poly.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Roots_of_Real_Complex_Poly.thy @@ -0,0 +1,245 @@ +section \Roots of Real and Complex Algebraic Polynomials\ + +text \We are now able to actually compute all roots of polynomials + with real and complex algebraic coefficients. The main addition to + calculating the representative polynomial for a superset of all roots + is to find the genuine roots. For this we utilize the approximation algorithm + via interval arithmetic.\ + +theory Roots_of_Real_Complex_Poly + imports + Roots_of_Algebraic_Poly_Impl + Roots_via_IA + MPoly_Container +begin + +hide_const (open) Module.smult + +typedef (overloaded) 'a rf_poly = "{ p :: 'a :: idom poly. rsquarefree p}" + by (intro exI[of _ 1], auto simp: rsquarefree_def) + +setup_lifting type_definition_rf_poly + +context +begin +lifting_forget poly.lifting + +lift_definition poly_rf :: "'a :: idom rf_poly \ 'a poly" is "\ x. x" . + +definition roots_of_poly_dummy :: "'a::{comm_ring_1,ring_no_zero_divisors} poly \ _" + where "roots_of_poly_dummy p = (SOME xs. set xs = {r. poly p r = 0} \ distinct xs)" + +lemma roots_of_poly_dummy_code[code]: + "roots_of_poly_dummy p = Code.abort (STR ''roots-of-poly-dummy'') (\ x. roots_of_poly_dummy p)" + by simp + +lemma roots_of_poly_dummy: assumes p: "p \ 0" + shows "set (roots_of_poly_dummy p) = {x. poly p x = 0}" "distinct (roots_of_poly_dummy p)" +proof - + from someI_ex[OF finite_distinct_list[OF poly_roots_finite[OF p]], folded roots_of_poly_dummy_def] + show "set (roots_of_poly_dummy p) = {x. poly p x = 0}" "distinct (roots_of_poly_dummy p)" by auto +qed + +lift_definition roots_of_complex_rf_poly_part1 :: "complex rf_poly \ complex genuine_roots_aux" is + "\ p. if Ball (set (Polynomial.coeffs p)) algebraic then + let q = representative_poly p; + zeros = complex_roots_of_int_poly q + in (p,zeros,Polynomial.degree p, filter_fun_complex p) + else (p,roots_of_poly_dummy p,Polynomial.degree p, filter_fun_complex p)" + subgoal for p + proof - + assume rp: "rsquarefree p" + hence card: "card {x. poly p x = 0} = Polynomial.degree p" + using rsquarefree_card_degree rsquarefree_def by blast + from rp have p: "p \ 0" unfolding rsquarefree_def by auto + have ff: "filter_fun p (filter_fun_complex p)" by (rule filter_fun_complex) + show ?thesis + proof (cases "Ball (set (Polynomial.coeffs p)) algebraic") + case False + with roots_of_poly_dummy[OF p] ff + show ?thesis using rp card by auto + next + case True + from rp card representative_poly_complex[of p] + complex_roots_of_int_poly[of "representative_poly p"] ff + show ?thesis unfolding Let_def rsquarefree_def using True by auto + qed + qed + done + + +lift_definition roots_of_real_rf_poly_part1 :: "real rf_poly \ real genuine_roots_aux" is + "\ p. let n = count_roots p in + if Ball (set (Polynomial.coeffs p)) algebraic then + let q = representative_poly p; + zeros = real_roots_of_int_poly q + in (p,zeros,n, filter_fun_real p) + else (p,roots_of_poly_dummy p,n, filter_fun_real p)" + subgoal for p + proof - + assume rp: "rsquarefree p" + from rp have p: "p \ 0" unfolding rsquarefree_def by auto + have ff: "filter_fun p (filter_fun_real p)" by (rule filter_fun_real) + show ?thesis + proof (cases "Ball (set (Polynomial.coeffs p)) algebraic") + case False + with roots_of_poly_dummy[OF p] ff + show ?thesis using rp by (auto simp: Let_def count_roots_correct) + next + case True + from rp representative_poly_real[of p] + real_roots_of_int_poly[of "representative_poly p"] ff + show ?thesis unfolding Let_def rsquarefree_def using True + by (auto simp: count_roots_correct) + qed + qed + done + +definition roots_of_complex_rf_poly :: "complex rf_poly \ complex list" where + "roots_of_complex_rf_poly p = genuine_roots_impl (roots_of_complex_rf_poly_part1 p)" + +lemma roots_of_complex_rf_poly: "set (roots_of_complex_rf_poly p) = {x. poly (poly_rf p) x = 0}" + "distinct (roots_of_complex_rf_poly p)" + unfolding roots_of_complex_rf_poly_def genuine_roots_impl + by (transfer, auto simp: genuine_roots_impl) + +definition roots_of_real_rf_poly :: "real rf_poly \ real list" where + "roots_of_real_rf_poly p = genuine_roots_impl (roots_of_real_rf_poly_part1 p)" + +lemma roots_of_real_rf_poly: "set (roots_of_real_rf_poly p) = {x. poly (poly_rf p) x = 0}" + "distinct (roots_of_real_rf_poly p)" + unfolding roots_of_real_rf_poly_def genuine_roots_impl + by (transfer, auto simp: genuine_roots_impl Let_def) + +typedef (overloaded) 'a rf_polys = "{ (a :: 'a :: idom, ps :: ('a poly \ nat) list). Ball (fst ` set ps) rsquarefree}" + by (intro exI[of _ "(_,Nil)"], auto) + +setup_lifting type_definition_rf_polys + +lift_definition yun_polys :: "'a :: {euclidean_ring_gcd,field_char_0,semiring_gcd_mult_normalize} poly \ 'a rf_polys" + is "\ p. yun_factorization gcd p" + subgoal for p + apply auto + apply (intro square_free_rsquarefree) + apply (insert yun_factorization[of p, OF refl]) + by (cases "yun_factorization gcd p", auto dest: square_free_factorizationD) + done + +context + notes [[typedef_overloaded]] +begin +lift_definition (code_dt) yun_rf :: "'a :: idom rf_polys \ 'a \ ('a rf_poly \ nat) list" is "\ x. x" + by (auto simp: list_all_iff, force) +end +end +definition polys_rf :: "'a :: idom rf_polys \ 'a rf_poly list" where + "polys_rf = map fst o snd o yun_rf" + +lemma yun_polys: assumes "p \ 0" + shows "poly p x = 0 \ (\ q \ set (polys_rf (yun_polys p)). poly (poly_rf q) x = 0)" + using assms unfolding polys_rf_def o_def + apply transfer + subgoal for p x + proof - + assume p: "p \ 0" + obtain c ps where yun: "yun_factorization gcd p = (c,ps)" by force + from yun_factorization[OF this] have sff: "square_free_factorization p (c, ps)" by auto + from square_free_factorizationD'(1)[OF sff] p have c0: "c \ 0" by auto + show ?thesis unfolding yun + unfolding square_free_factorizationD'(1)[OF sff] poly_smult poly_prod_list snd_conv + mult_eq_0_iff prod_list_zero_iff + using c0 by force + qed + done + + +definition roots_of_complex_rf_polys :: "complex rf_polys \ complex list" where + "roots_of_complex_rf_polys ps = concat (map roots_of_complex_rf_poly (polys_rf ps))" + +lemma roots_of_complex_rf_polys: + "set (roots_of_complex_rf_polys ps) = {x. \ p \ set (polys_rf ps). poly (poly_rf p) x = 0 }" + unfolding roots_of_complex_rf_polys_def set_concat set_map image_comp o_def + roots_of_complex_rf_poly by auto + +definition roots_of_real_rf_polys :: "real rf_polys \ real list" where + "roots_of_real_rf_polys ps = concat (map roots_of_real_rf_poly (polys_rf ps))" + +lemma roots_of_real_rf_polys: + "set (roots_of_real_rf_polys ps) = {x. \ p \ set (polys_rf ps). poly (poly_rf p) x = 0 }" + unfolding roots_of_real_rf_polys_def set_concat set_map image_comp o_def + roots_of_real_rf_poly by auto + +definition roots_of_complex_poly :: "complex poly \ complex list" where + "roots_of_complex_poly p = (if p = 0 then [] else roots_of_complex_rf_polys (yun_polys p))" + +lemma roots_of_complex_poly: assumes p: "p \ 0" + shows "set (roots_of_complex_poly p) = {x. poly p x = 0}" + using p unfolding roots_of_complex_poly_def + by (simp add: roots_of_complex_rf_polys yun_polys[OF p]) + +definition roots_of_real_poly :: "real poly \ real list" where + "roots_of_real_poly p = (if p = 0 then [] else roots_of_real_rf_polys (yun_polys p))" + +lemma roots_of_real_poly: assumes p: "p \ 0" + shows "set (roots_of_real_poly p) = {x. poly p x = 0}" + using p unfolding roots_of_real_poly_def + by (simp add: roots_of_real_rf_polys yun_polys[OF p]) + +lemma distinct_concat': + "\ distinct (list_neq xs []); + \ ys. ys \ set xs \ distinct ys; + \ ys zs. \ ys \ set xs ; zs \ set xs ; ys \ zs \ \ set ys \ set zs = {} + \ \ distinct (concat xs)" + by (induct xs, auto split: if_splits) + +lemma roots_of_rf_yun_polys_distinct: assumes + rt: "\ p. set (rop p) = {x. poly (poly_rf p) x = 0}" + and dist: "\ p. distinct (rop p)" +shows "distinct (concat (map rop (polys_rf (yun_polys p))))" + using assms unfolding polys_rf_def +proof (transfer, goal_cases) + case (1 rop p) + obtain c fs where yun: "yun_factorization gcd p = (c,fs)" by force + note sff = yun_factorization(1)[OF yun] + note sff1 = square_free_factorizationD[OF sff] + note sff2 = square_free_factorizationD'[OF sff] + have rs: "(p,i) \ set fs \ rsquarefree p" for p i + by (intro square_free_rsquarefree, insert sff1(2), auto) + note 1 = 1[OF rs] + show ?case unfolding yun snd_conv map_map o_def using 1 sff1(3,5) + proof (induct fs) + case (Cons pi fs) + obtain p i where pi: "pi = (p,i)" by force + hence "(p,i) \ set (pi # fs)" by auto + note p_i = Cons(2-4)[OF this] + have IH: "distinct (concat (map (\x. rop (fst x)) fs))" + by (rule Cons(1)[OF Cons(2,3,4)], insert Cons(5), auto) + { + fix x + assume x: "x \ set (rop p)" "x \ (\x\set fs. set (rop (fst x)))" + from x[unfolded p_i] have rtp: "poly p x = 0" by auto + from x obtain q j where qj: "(q,j) \ set fs" and x: "x \ set (rop q)" by force + from Cons(2)[of q j] x qj have rtq: "poly q x = 0" by auto + from Cons(5)[unfolded pi] qj have "(p,i) \ (q,j)" by auto + from p_i(3)[OF _ this] qj have cop: "algebraic_semidom_class.coprime p q" by auto + from rtp have dvdp: "[:-x,1:] dvd p" using poly_eq_0_iff_dvd by blast + from rtq have dvdq: "[:-x,1:] dvd q" using poly_eq_0_iff_dvd by blast + from cop dvdp dvdq have "is_unit [:-x,1:]" by (metis coprime_common_divisor) + hence False by simp + } + thus ?case unfolding pi by (auto simp: p_i(2) IH) + qed simp +qed + +lemma distinct_roots_of_real_poly: "distinct (roots_of_real_poly p)" + unfolding roots_of_real_poly_def roots_of_real_rf_polys_def + using roots_of_rf_yun_polys_distinct[of roots_of_real_rf_poly p, OF roots_of_real_rf_poly] + by auto + +lemma distinct_roots_of_complex_poly: "distinct (roots_of_complex_poly p)" + unfolding roots_of_complex_poly_def roots_of_complex_rf_polys_def + using roots_of_rf_yun_polys_distinct[of roots_of_complex_rf_poly p, OF roots_of_complex_rf_poly] + by auto + + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/Roots_via_IA.thy b/thys/Factor_Algebraic_Polynomial/Roots_via_IA.thy new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/Roots_via_IA.thy @@ -0,0 +1,303 @@ +section \Root Filter via Interval Arithmetic\ + +subsection \Generic Framework\ + +text \We provide algorithms for finding all real or complex roots of a polynomial + from a superset of the roots via interval arithmetic. + These algorithms are much faster than just + evaluating the polynomial via algebraic number computations.\ + +theory Roots_via_IA + imports + Algebraic_Numbers.Interval_Arithmetic +begin + +definition interval_of_real :: "nat \ real \ real interval" where + "interval_of_real prec x = + (if is_rat x then Interval x x + else let n = 2 ^ prec; x' = x * of_int n + in Interval (of_rat (Rat.Fract \x'\ n)) (of_rat (Rat.Fract \x'\ n)))" + +definition interval_of_complex :: "nat \ complex \ complex_interval" where + "interval_of_complex prec z = + Complex_Interval (interval_of_real prec (Re z)) (interval_of_real prec (Im z))" + +fun poly_interval :: "'a :: {plus,times,zero} list \ 'a \ 'a" where + "poly_interval [] _ = 0" +| "poly_interval [c] _ = c" +| "poly_interval (c # cs) x = c + x * poly_interval cs x" + +definition filter_fun_complex :: "complex poly \ nat \ complex \ bool" where + "filter_fun_complex p = (let c = coeffs p in + (\ prec. let cs = map (interval_of_complex prec) c + in (\ x. 0 \\<^sub>c poly_interval cs (interval_of_complex prec x))))" + +definition filter_fun_real :: "real poly \ nat \ real \ bool" where + "filter_fun_real p = (let c = coeffs p in + (\ prec. let cs = map (interval_of_real prec) c + in (\ x. 0 \\<^sub>i poly_interval cs (interval_of_real prec x))))" + +definition genuine_roots :: "_ poly \ _ list \ _ list" where + "genuine_roots p xs = filter (\x. poly p x = 0) xs" + +lemma zero_in_interval_0 [simp, intro]: "0 \\<^sub>i 0" + unfolding zero_interval_def by auto + +lemma zero_in_complex_interval_0 [simp, intro]: "0 \\<^sub>c 0" + unfolding zero_complex_interval_def by (auto simp: in_complex_interval_def) + +lemma length_coeffs_degree': + "length (coeffs p) = (if p = 0 then 0 else Suc (degree p))" + by (cases "p = 0") (auto simp: length_coeffs_degree) + +lemma poly_in_poly_interval_complex: + assumes "list_all2 (\c ivl. c \\<^sub>c ivl) (coeffs p) cs" "x \\<^sub>c ivl" + shows "poly p x \\<^sub>c poly_interval cs ivl" +proof - + have len_eq: "length (coeffs p) = length cs" + using assms(1) list_all2_lengthD by blast + have "coeffs p = map (\i. coeffs p ! i) [0.. = map (poly.coeff p) [0..c ivl. c \\<^sub>c ivl) (map (poly.coeff p) [0.. length (coeffs p)" + using len_eq by simp + ultimately show ?thesis using assms(2) + proof (induction cs ivl arbitrary: p x rule: poly_interval.induct) + case (1 ivl p x) + thus ?case by auto + next + case (2 c ivl p x) + have "degree p = 0" + using 2 by (auto simp: degree_eq_length_coeffs) + then obtain c' where [simp]: "p = [:c':]" + by (meson degree_eq_zeroE) + show ?case using 2 by auto + next + case (3 c1 c2 cs ivl p x) + obtain q c where [simp]: "p = pCons c q" + by (cases p rule: pCons_cases) + have "list_all2 in_complex_interval (map (poly.coeff p) [0.. = c # map (poly.coeff q) [0..\<^sub>c c1" and + "list_all2 in_complex_interval (map (poly.coeff q) [0..\<^sub>c poly_interval (c2 # cs) ivl" + proof (rule "3.IH") + show "length (coeffs q) \ length (c2 # cs)" + using "3.prems"(2) unfolding length_coeffs_degree' by auto + qed fact+ + + show ?case + using IH "3.prems" \c \\<^sub>c c1\ + by (auto intro!: plus_complex_interval times_complex_interval) + qed +qed + +lemma poly_in_poly_interval_real: fixes x :: real + assumes "list_all2 (\c ivl. c \\<^sub>i ivl) (coeffs p) cs" "x \\<^sub>i ivl" + shows "poly p x \\<^sub>i poly_interval cs ivl" +proof - + have len_eq: "length (coeffs p) = length cs" + using assms(1) list_all2_lengthD by blast + have "coeffs p = map (\i. coeffs p ! i) [0.. = map (poly.coeff p) [0..c ivl. c \\<^sub>i ivl) (map (poly.coeff p) [0.. length (coeffs p)" + using len_eq by simp + ultimately show ?thesis using assms(2) + proof (induction cs ivl arbitrary: p x rule: poly_interval.induct) + case (1 ivl p x) + thus ?case by auto + next + case (2 c ivl p x) + have "degree p = 0" + using 2 by (auto simp: degree_eq_length_coeffs) + then obtain c' where [simp]: "p = [:c':]" + by (meson degree_eq_zeroE) + show ?case using 2 by auto + next + case (3 c1 c2 cs ivl p x) + obtain q c where [simp]: "p = pCons c q" + by (cases p rule: pCons_cases) + have "list_all2 in_interval (map (poly.coeff p) [0.. = c # map (poly.coeff q) [0..\<^sub>i c1" and + "list_all2 in_interval (map (poly.coeff q) [0..\<^sub>i poly_interval (c2 # cs) ivl" + proof (rule "3.IH") + show "length (coeffs q) \ length (c2 # cs)" + using "3.prems"(2) unfolding length_coeffs_degree' by auto + qed fact+ + + show ?case + using IH "3.prems" \c \\<^sub>i c1\ + by (auto intro!: plus_in_interval times_in_interval) + qed +qed + + +lemma in_interval_of_real [simp, intro]: "x \\<^sub>i interval_of_real prec x" + unfolding interval_of_real_def by (auto simp: Let_def of_rat_rat field_simps) + +lemma in_interval_of_complex [simp, intro]: "z \\<^sub>c interval_of_complex prec z" + unfolding interval_of_complex_def in_complex_interval_def by auto + +lemma distinct_genuine_roots [simp, intro]: + "distinct xs \ distinct (genuine_roots p xs)" + by (simp add: genuine_roots_def) + +definition filter_fun :: "'a poly \ (nat \ 'a :: comm_ring \ bool) \ bool" where + "filter_fun p f = (\ n x. poly p x = 0 \ f n x)" + +lemma filter_fun_complex: "filter_fun p (filter_fun_complex p)" + unfolding filter_fun_def +proof (intro impI allI) + fix prec x + assume root: "poly p x = 0" + define cs where "cs = map (interval_of_complex prec) (coeffs p)" + have cs: "list_all2 in_complex_interval (coeffs p) cs" + unfolding cs_def list_all2_map2 by (intro list_all2_refl in_interval_of_complex) + define P where "P = (\x. 0 \\<^sub>c poly_interval cs (interval_of_complex prec x))" + have "P x" + proof - + have "poly p x \\<^sub>c poly_interval cs (interval_of_complex prec x)" + by (intro poly_in_poly_interval_complex in_interval_of_complex cs) + with root show ?thesis + by (simp add: P_def) + qed + thus "filter_fun_complex p prec x" unfolding filter_fun_complex_def Let_def P_def + using cs_def by blast +qed + +lemma filter_fun_real: "filter_fun p (filter_fun_real p)" + unfolding filter_fun_def +proof (intro impI allI) + fix prec x + assume root: "poly p x = 0" + define cs where "cs = map (interval_of_real prec) (coeffs p)" + have cs: "list_all2 in_interval (coeffs p) cs" + unfolding cs_def list_all2_map2 by (intro list_all2_refl in_interval_of_real) + define P where "P = (\x. 0 \\<^sub>i poly_interval cs (interval_of_real prec x))" + have "P x" + proof - + have "poly p x \\<^sub>i poly_interval cs (interval_of_real prec x)" + by (intro poly_in_poly_interval_real in_interval_of_real cs) + with root show ?thesis + by (simp add: P_def) + qed + thus "filter_fun_real p prec x" unfolding filter_fun_real_def Let_def P_def + using cs_def by blast +qed + +context + fixes p :: "'a :: comm_ring poly" and f + assumes ff: "filter_fun p f" +begin + +lemma genuine_roots_step: + "genuine_roots p xs = genuine_roots p (filter (f prec) xs)" + unfolding genuine_roots_def filter_filter + using ff[unfolded filter_fun_def, rule_format, of _ prec] by metis + +lemma genuine_roots_step_preserve_invar: + assumes "{z. poly p z = 0} \ set xs" + shows "{z. poly p z = 0} \ set (filter (f prec) xs)" +proof - + have "{z. poly p z = 0} = set (genuine_roots p xs)" + using assms by (auto simp: genuine_roots_def) + also have "\ = set (genuine_roots p (filter (f prec) xs))" + using genuine_roots_step[of _ prec] by simp + also have "\ \ set (filter (f prec) xs)" + by (auto simp: genuine_roots_def) + finally show ?thesis . +qed +end + +lemma genuine_roots_finish: + fixes p :: "'a :: field_char_0 poly" + assumes "{z. poly p z = 0} \ set xs" "distinct xs" + assumes "length xs = card {z. poly p z = 0}" + shows "genuine_roots p xs = xs" +proof - + have [simp]: "p \ 0" + using finite_subset[OF assms(1) finite_set] infinite_UNIV_char_0 by auto + have "length (genuine_roots p xs) = length xs" + unfolding genuine_roots_def using assms + by (simp add: Int_absorb2 distinct_length_filter) + thus ?thesis + unfolding genuine_roots_def + by (metis filter_True length_filter_less linorder_not_less order_eq_iff) +qed + +text \This is type of the initial search problem. It consists of a polynomial $p$, + a list $xs$ of candidate roots, the cardinality of the set of roots of $p$ and a filter function to + drop non-roots that is parametric in a precision parameter.\ +typedef (overloaded) 'a genuine_roots_aux = + "{(p :: 'a :: field_char_0 poly, xs, n, ff). + distinct xs \ + {z. poly p z = 0} \ set xs \ + card {z. poly p z = 0} = n \ + filter_fun p ff}" + by (rule exI[of _ "(1, [], 0, \ _ _. False)"], auto simp: filter_fun_def) + +setup_lifting type_definition_genuine_roots_aux + +lift_definition genuine_roots' :: "nat \ 'a :: field_char_0 genuine_roots_aux \ 'a list" is + "\prec (p, xs, n, ff). genuine_roots p xs" . + +lift_definition genuine_roots_impl_step' :: "nat \ 'a :: field_char_0 genuine_roots_aux \ 'a genuine_roots_aux" is + "\prec (p, xs, n, ff). (p, filter (ff prec) xs, n, ff)" + by (safe, intro distinct_filter, auto simp: filter_fun_def) + +lift_definition gr_poly :: "'a :: field_char_0 genuine_roots_aux \ 'a poly" is + "\(p :: 'a poly, _, _, _). p" . + +lift_definition gr_list :: "'a :: field_char_0 genuine_roots_aux \ 'a list" is + "\(_, xs :: 'a list, _, _). xs" . + +lift_definition gr_numroots :: "'a :: field_char_0 genuine_roots_aux \ nat" is + "\(_, _, n, _). n" . + +lemma genuine_roots'_code [code]: + "genuine_roots' prec gr = + (if length (gr_list gr) = gr_numroots gr then gr_list gr + else genuine_roots' (2 * prec) (genuine_roots_impl_step' prec gr))" +proof (transfer, clarify) + fix prec :: nat and p :: "'a poly" and xs :: "'a list" and ff + assume *: "{z. poly p z = 0} \ set xs" "distinct xs" "filter_fun p ff" + show "genuine_roots p xs = + (if length xs = card {z. poly p z = 0} then xs + else genuine_roots p (filter (ff prec) xs))" + using genuine_roots_finish[of p xs] genuine_roots_step[of p] * by auto +qed + +definition initial_precision :: nat where "initial_precision = 10" + +definition genuine_roots_impl :: "'a genuine_roots_aux \ 'a :: field_char_0 list" where + "genuine_roots_impl = genuine_roots' initial_precision" + +lemma genuine_roots_impl: "set (genuine_roots_impl p) = {z. poly (gr_poly p) z = 0}" + "distinct (genuine_roots_impl p)" + unfolding genuine_roots_impl_def + by (transfer, auto simp: genuine_roots_def, transfer, auto) + +end \ No newline at end of file diff --git a/thys/Factor_Algebraic_Polynomial/document/root.bib b/thys/Factor_Algebraic_Polynomial/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/document/root.bib @@ -0,0 +1,64 @@ +@article{Algebraic_Numbers-AFP, + author = {René Thiemann and Akihisa Yamada and Sebastiaan Joosten}, + title = {Algebraic Numbers in {I}sabelle/{HOL}}, + journal = {Archive of Formal Proofs}, + month = dec, + year = 2015, + note = {\url{https://isa-afp.org/entries/Algebraic_Numbers.html}, + Formal proof development}, + ISSN = {2150-914x}, +} + +@article{Algebraic_Numbers, + author = {Sebastiaan J. C. Joosten and + Ren{\'{e}} Thiemann and + Akihisa Yamada}, + title = {A Verified Implementation of Algebraic Numbers in {I}sabelle/{HOL}}, + journal = {J. Autom. Reason.}, + volume = {64}, + number = {3}, + pages = {363--389}, + year = {2020} +} + +@article{jsc, + title = {Computing in the Field of Complex Algebraic Numbers}, + author = {Adam Wojciech Strzebo\'nski}, + year = 1997, + volume = 24, + pages = {647-656}, + journal = {J.\ Symbolic Computation}, +} + +@article{Brown, + author = {Brown, W. S.}, + title = {The Subresultant {PRS} Algorithm}, + journal = {ACM Trans. Math. Softw.}, + volume = {4}, + number = {3}, + year = {1978}, + pages = {237--249}, + doi = {10.1145/355791.355795}, +} + +@article{BrownTraub, + author = {Brown, W. S. and Traub, J. F.}, + title = {On {E}uclid's Algorithm and the Theory of Subresultants}, + journal = {Journal of the ACM}, + volume = {18}, + number = {4}, + year = {1971}, + pages = {505--514}, + doi = {10.1145/321662.321665}, +} + +@article{Subresultants-AFP, + author = {Sebastiaan Joosten and René Thiemann and Akihisa Yamada}, + title = {Subresultants}, + journal = {Archive of Formal Proofs}, + month = apr, + year = 2017, + note = {\url{https://isa-afp.org/entries/Subresultants.html}, + Formal proof development}, + ISSN = {2150-914x}, +} diff --git a/thys/Factor_Algebraic_Polynomial/document/root.tex b/thys/Factor_Algebraic_Polynomial/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Factor_Algebraic_Polynomial/document/root.tex @@ -0,0 +1,75 @@ +\documentclass[11pt,a4paper]{article} + +\usepackage{isabelle,isabellesym} +\usepackage{url} + +\usepackage{amssymb} +\usepackage{amsmath} +\usepackage{amsthm} + +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\newcommand\rats{\mathbb{Q}} +\newcommand\ints{\mathbb{Z}} +\newcommand\reals{\mathbb{R}} +\newcommand\complex{\mathbb{C}} + +\begin{document} + +\title{Factorization of Polynomials with Algebraic Coefficients\footnote{Supported by FWF (Austrian Science Fund) project Y757.}} +\author{Manuel Eberl \and Ren\'e Thiemann} +\maketitle + +\begin{abstract} +The AFP already contains a verified implementation of algebraic numbers. +However, it is has a severe limitation in its +factorization algorithm of real and complex polynomials: the factorization is only +guaranteed to succeed if the coefficients of the polynomial are rational numbers. +In this work, we verify an algorithm to factor all real and complex polynomials +whose coefficients are algebraic. +The existence of such an algorithm proves in a constructive way that the set of complex algebraic numbers +is algebraically closed. +Internally, the algorithm is based on resultants of multivariate +polynomials and an approximation algorithm using interval arithmetic. +\end{abstract} + +\tableofcontents + +\section{Introduction} + +The formalization of algebraic numbers \cite{Algebraic_Numbers,Algebraic_Numbers-AFP} +includes an algorithm that given +a univariate polynomial $f$ over $\ints$ or $\rats$, it computes all roots of $f$ within +$\reals$ or $\complex$. In this AFP entry we verify a generalized algorithm that also allows +polynomials as input whose coefficients are complex or real algebraic numbers, following +\cite[Section~3]{jsc}. + +The verified algorithm internally computes resultants of +multivariate polynomials, where we utilize Braun and Traub's subresultant algorithm in our +verified implementation \cite{Brown,BrownTraub,Subresultants-AFP}. In this way we achieve an +efficient implementation with minimal effort: only a division algorithm for multivariate polynomials +is required, but no algorithm for computing greatest common divisors of these polynomials. + +\paragraph{Acknowledgments} +We thank Dmitriy Traytel for help with code generation for functions defined via +\isacommand{lift{\isacharunderscore}{\kern0pt}definition}. + + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/Linear_Recurrences/ROOT b/thys/Linear_Recurrences/ROOT --- a/thys/Linear_Recurrences/ROOT +++ b/thys/Linear_Recurrences/ROOT @@ -1,25 +1,26 @@ chapter AFP session Linear_Recurrences (AFP) = Count_Complex_Roots + options [timeout = 600] sessions "HOL-Library" "HOL-Combinatorics" "HOL-Computational_Algebra" "HOL-Analysis" "HOL-Real_Asymp" Polynomial_Factorization theories Linear_Homogenous_Recurrences Linear_Inhomogenous_Recurrences Rational_FPS_Asymptotics document_files "root.tex" session Linear_Recurrences_Solver (AFP) in "Solver" = Berlekamp_Zassenhaus + options [timeout = 1200] sessions Algebraic_Numbers + Factor_Algebraic_Polynomial Linear_Recurrences theories Linear_Recurrences_Test diff --git a/thys/Linear_Recurrences/Solver/Linear_Recurrences_Solver.thy b/thys/Linear_Recurrences/Solver/Linear_Recurrences_Solver.thy --- a/thys/Linear_Recurrences/Solver/Linear_Recurrences_Solver.thy +++ b/thys/Linear_Recurrences/Solver/Linear_Recurrences_Solver.thy @@ -1,106 +1,107 @@ (* File: Linear_Recurrences_Solver.thy Author: Manuel Eberl, TU München *) section \Solver for linear recurrences\ theory Linear_Recurrences_Solver imports Complex_Main Linear_Recurrences.Linear_Homogenous_Recurrences Linear_Recurrences.Linear_Inhomogenous_Recurrences - "Algebraic_Numbers.Complex_Algebraic_Numbers" + Factor_Algebraic_Polynomial.Factor_Complex_Poly begin -lemma factorize_complex_main: - assumes "factorize_complex_main p = Some fctrs" +lemma is_factorization_of_factor_complex_main: + assumes "factor_complex_main p = fctrs" shows "is_factorization_of fctrs p" unfolding is_factorization_of_def proof safe from assms have "p = Polynomial.smult (fst fctrs) (\(x, i)\snd fctrs. [:- x, 1:] ^ Suc i)" - by (intro factorize_complex_main) simp_all + by (intro factor_complex_main) simp_all also have "\ = interp_factorization fctrs" by (simp add: interp_factorization_def case_prod_unfold) finally show "interp_factorization fctrs = p" .. -qed (insert assms, simp_all add: distinct_factorize_complex_main) + show "distinct (map fst (snd fctrs))" unfolding assms[symmetric] + by (rule distinct_factor_complex_main) +qed definition solve_ratfps - :: "complex ratfps \ (complex poly \ (complex poly \ complex) list) option" where + :: "complex ratfps \ complex poly \ (complex poly \ complex) list" where "solve_ratfps f = (case quot_of_ratfps f of (p, q) \ - map_option (solve_factored_ratfps' p) (factorize_complex_main (reflect_poly q)))" + solve_factored_ratfps' p (factor_complex_main (reflect_poly q)))" lemma solve_ratfps: - assumes "solve_ratfps f = Some sol" + assumes "solve_ratfps f = sol" shows "Abs_fps (interp_ratfps_solution sol) = fps_of_ratfps f" proof - define p and q where "p = fst (quot_of_ratfps f)" and "q = snd (quot_of_ratfps f)" - with assms obtain fctrs where fctrs: "factorize_complex_main (reflect_poly q) = Some fctrs" + with assms obtain fctrs where fctrs: "factor_complex_main (reflect_poly q) = fctrs" by (auto simp: solve_ratfps_def p_def q_def case_prod_unfold split: if_splits) have q: "coeff q 0 \ 0" by (simp add: q_def) hence [simp]: "q \ 0" by auto from fctrs have "is_factorization_of fctrs (reflect_poly q)" - by (rule factorize_complex_main) + by (rule is_factorization_of_factor_complex_main) with assms have "is_alt_factorization_of fctrs (reflect_poly (reflect_poly q))" by (intro reflect_factorization) simp_all hence "is_alt_factorization_of fctrs q" by (simp add: q) with fctrs q have "Abs_fps (interp_ratfps_solution (solve_factored_ratfps' p fctrs)) = fps_of_poly p / fps_of_poly q" by (intro solve_factored_ratfps') (simp_all) also from fctrs assms have "solve_factored_ratfps' p fctrs = sol" by (simp add: solve_ratfps_def p_def q_def case_prod_unfold split: if_splits) finally show ?thesis by (simp add: fps_of_ratfps_altdef case_prod_unfold p_def q_def) qed definition solve_lhr :: "complex list \ complex list \ (complex poly \ (complex poly \ complex) list) option" where "solve_lhr cs fs = (if cs = [] \ length fs < length cs - 1 then None else let m = length fs + 1 - length cs; p = lhr_fps_numerator m cs (\n. fs ! n); q = lr_fps_denominator' cs - in map_option (solve_factored_ratfps' p) (factorize_complex_main q))" + in Some (solve_factored_ratfps' p (factor_complex_main q)))" lemma solve_lhr: assumes "linear_homogenous_recurrence f cs fs" assumes "Some sol = solve_lhr cs fs" shows "f = interp_ratfps_solution sol" proof - - from assms(2) obtain fctrs where - fctrs: "factorize_complex_main (lr_fps_denominator' cs) = Some fctrs" - by (cases "factorize_complex_main (lr_fps_denominator' cs)") - (simp_all add: solve_lhr_def Let_def split: if_splits) - from factorize_complex_main[OF this] + obtain fctrs where + fctrs: "factor_complex_main (lr_fps_denominator' cs) = fctrs" + by auto + from is_factorization_of_factor_complex_main[OF this] have factorization: "is_factorization_of fctrs (lr_fps_denominator' cs)" . have "f = interp_ratfps_solution (solve_factored_ratfps' (lhr_fps_numerator (length fs + 1 - length cs) cs ((!) fs)) fctrs)" (is "_ = interp_ratfps_solution ?sol") by (intro solve_lhr_aux) fact+ also from assms(2) have "?sol = sol" by (auto simp: solve_lhr_def Let_def case_prod_unfold fctrs split: if_splits) finally show ?thesis . qed definition solve_lir :: "complex list \ complex list \ complex polyexp \ (complex poly \ (complex poly \ complex) list) option" where - "solve_lir cs fs g = Option.bind (lir_fps cs fs g) solve_ratfps" + "solve_lir cs fs g = map_option solve_ratfps (lir_fps cs fs g)" lemma solve_lir: assumes "linear_inhomogenous_recurrence f (eval_polyexp g) cs fs" assumes "solve_lir cs fs g = Some sol" shows "f = interp_ratfps_solution sol" proof - from lir_fps_correct[OF assms(1)] obtain fps where fps: "lir_fps cs fs g = Some fps" "fps_of_ratfps fps = Abs_fps f" by blast - from assms(2) have "solve_ratfps fps = Some sol" + from assms(2) have "solve_ratfps fps = sol" by (simp add: solve_lir_def fps case_prod_unfold) from solve_ratfps[OF this] have "Abs_fps (interp_ratfps_solution sol) = fps_of_ratfps fps" by (simp add: case_prod_unfold fps_of_ratfps_altdef) with fps have "Abs_fps f = Abs_fps (interp_ratfps_solution sol)" by simp thus ?thesis by (simp add: fun_eq_iff fps_eq_iff) qed end diff --git a/thys/PAL/PAL.thy b/thys/PAL/PAL.thy new file mode 100644 --- /dev/null +++ b/thys/PAL/PAL.thy @@ -0,0 +1,158 @@ +section \Public Announcement Logic (PAL) in HOL\ + +text \An earlier encoding and automation of the wise men puzzle, utilizing a shallow embedding of +higher-order (multi-)modal logic in HOL, has been presented in \cite{J41,J44}. However, this work did not +convincingly address the interaction dynamics between the involved agents. Here we therefore extend and adapt +the universal (meta-)logical reasoning approach of \cite{J41} for public announcement logic (PAL) and +we demonstrate how it can be utilized to achieve a convincing encoding and automation of the +wise men puzzle in HOL, so that also the interaction dynamics as given in the scenario is adequately +addressed. For further background information on the work presented here we refer to \cite{R78,C90}.\ + +theory PAL imports Main begin (* Sebastian Reiche and Christoph Benzmüller, 2021 *) +nitpick_params[user_axioms,expect=genuine] + +text \Type i is associated with possible worlds\ + typedecl i (* Type of possible worlds *) + type_synonym \ = "i\bool" (*Type of world domains *) + type_synonym \ = "\\i\bool" (* Type of world depended formulas (truth sets) *) + type_synonym \ = "i\i\bool" (* Type of accessibility relations between world *) + type_synonym \ = "\\bool" (* Type of groups of agents *) + +text \Some useful relations (for constraining accessibility relations)\ +definition reflexive::"\\bool" + where "reflexive R \ \x. R x x" +definition symmetric::"\\bool" + where "symmetric R \ \x y. R x y \ R y x" +definition transitive::"\\bool" + where "transitive R \ \x y z. R x y \ R y z \ R x z" +definition euclidean::"\\bool" + where "euclidean R \ \x y z. R x y \ R x z \ R y z" +definition intersection_rel::"\\\\\" + where "intersection_rel R Q \ \u v. R u v \ Q u v" +definition union_rel::"\\\\\" + where "union_rel R Q \ \u v. R u v \ Q u v" +definition sub_rel::"\\\\bool" + where "sub_rel R Q \ \u v. R u v \ Q u v" +definition inverse_rel::"\\\" + where "inverse_rel R \ \u v. R v u" +definition big_union_rel::"\\\" + where "big_union_rel X \ \u v. \R. (X R) \ (R u v)" +definition big_intersection_rel::"\\\" + where "big_intersection_rel X \ \u v. \R. (X R) \ (R u v)" + +text \In HOL the transitive closure of a relation can be defined in a single line.\ +definition tc::"\\\" + where "tc R \ \x y.\Q. transitive Q \ (sub_rel R Q \ Q x y)" + +text \Logical connectives for PAL\ +abbreviation patom::"\\\" ("\<^sup>A_"[79]80) + where "\<^sup>Ap \ \W w. W w \ p w" +abbreviation ptop::"\" ("\<^bold>\") + where "\<^bold>\ \ \W w. True" +abbreviation pneg::"\\\" ("\<^bold>\_"[52]53) + where "\<^bold>\\ \ \W w. \(\ W w)" +abbreviation pand::"\\\\\" (infixr"\<^bold>\"51) + where "\\<^bold>\\ \ \W w. (\ W w) \ (\ W w)" +abbreviation por::"\\\\\" (infixr"\<^bold>\"50) + where "\\<^bold>\\ \ \W w. (\ W w) \ (\ W w)" +abbreviation pimp::"\\\\\" (infixr"\<^bold>\"49) + where "\\<^bold>\\ \ \W w. (\ W w) \ (\ W w)" +abbreviation pequ::"\\\\\" (infixr"\<^bold>\"48) + where "\\<^bold>\\ \ \W w. (\ W w) \ (\ W w)" +abbreviation pknow::"\\\\\" ("\<^bold>K_ _") + where "\<^bold>K r \ \ \W w.\v. (W v \ r w v) \ (\ W v)" +abbreviation ppal::"\\\\\" ("\<^bold>[\<^bold>!_\<^bold>]_") + where "\<^bold>[\<^bold>!\\<^bold>]\ \ \W w. (\ W w) \ (\ (\z. W z \ \ W z) w)" + +text \Glogal validity of PAL formulas\ +abbreviation pvalid::"\ \ bool" ("\<^bold>\_\<^bold>\"[7]8) + where "\<^bold>\\\<^bold>\ \ \W.\w. W w \ \ W w" + +text \Introducing agent knowledge (K), mutual knowledge (E), distributed knowledge (D) and common knowledge (C).\ +abbreviation EVR::"\\\" + where "EVR G \ big_union_rel G" +abbreviation DIS::"\\\" + where "DIS G \ big_intersection_rel G" +abbreviation agttknows::"\\\\\" ("\<^bold>K\<^sub>_ _") + where "\<^bold>K\<^sub>r \ \ \<^bold>K r \" +abbreviation evrknows::"\\\\\" ("\<^bold>E\<^sub>_ _") + where "\<^bold>E\<^sub>G \ \ \<^bold>K (EVR G) \" +abbreviation disknows :: "\\\\\" ("\<^bold>D\<^sub>_ _") + where "\<^bold>D\<^sub>G \ \ \<^bold>K (DIS G) \" +abbreviation prck::"\\\\\\\" ("\<^bold>C\<^sub>_\<^bold>\_\<^bold>|_\<^bold>\") + where "\<^bold>C\<^sub>G\<^bold>\\\<^bold>|\\<^bold>\ \ \W w. \v. (tc (intersection_rel (EVR G) (\u v. W v \ \ W v)) w v) \ (\ W v)" +abbreviation pcmn::"\\\\\" ("\<^bold>C\<^sub>_ _") + where "\<^bold>C\<^sub>G \ \ \<^bold>C\<^sub>G\<^bold>\\<^bold>\\<^bold>|\\<^bold>\" + +text \Postulating S5 principles for the agent’s accessibility relations.\ +abbreviation S5Agent::"\\bool" + where "S5Agent i \ reflexive i \ transitive i \ euclidean i" +abbreviation S5Agents::"\\bool" + where "S5Agents A \ \i. (A i \ S5Agent i)" + +text \Introducing "Defs" as the set of the above definitions; useful for convenient unfolding.\ +named_theorems Defs +declare reflexive_def[Defs] symmetric_def[Defs] transitive_def[Defs] + euclidean_def[Defs] intersection_rel_def[Defs] union_rel_def[Defs] + sub_rel_def[Defs] inverse_rel_def[Defs] big_union_rel_def[Defs] + big_intersection_rel_def[Defs] tc_def[Defs] + +text \Consistency: nitpick reports a model.\ + lemma True nitpick [satisfy] oops (* model found *) + + +section \Automating the Wise Men Puzzle\ + +text \Agents are modeled as accessibility relations.\ +consts a::"\" b::"\" c::"\" +abbreviation Agent::"\\bool" ("\") where "\ x \ x = a \ x = b \ x = c" +axiomatization where group_S5: "S5Agents \" + +text \Common knowledge: At least one of a, b and c has a white spot.\ +consts ws::"\\\" +axiomatization where WM1: "\<^bold>\\<^bold>C\<^sub>\ (\<^sup>Aws a \<^bold>\ \<^sup>Aws b \<^bold>\ \<^sup>Aws c)\<^bold>\" + +text \Common knowledge: If x does not have a white spot then y knows this.\ +axiomatization where + WM2ab: "\<^bold>\\<^bold>C\<^sub>\ (\<^bold>\(\<^sup>Aws a) \<^bold>\ (\<^bold>K\<^sub>b (\<^bold>\(\<^sup>Aws a))))\<^bold>\" and + WM2ac: "\<^bold>\\<^bold>C\<^sub>\ (\<^bold>\(\<^sup>Aws a) \<^bold>\ (\<^bold>K\<^sub>c (\<^bold>\(\<^sup>Aws a))))\<^bold>\" and + WM2ba: "\<^bold>\\<^bold>C\<^sub>\ (\<^bold>\(\<^sup>Aws b) \<^bold>\ (\<^bold>K\<^sub>a (\<^bold>\(\<^sup>Aws b))))\<^bold>\" and + WM2bc: "\<^bold>\\<^bold>C\<^sub>\ (\<^bold>\(\<^sup>Aws b) \<^bold>\ (\<^bold>K\<^sub>c (\<^bold>\(\<^sup>Aws b))))\<^bold>\" and + WM2ca: "\<^bold>\\<^bold>C\<^sub>\ (\<^bold>\(\<^sup>Aws c) \<^bold>\ (\<^bold>K\<^sub>a (\<^bold>\(\<^sup>Aws c))))\<^bold>\" and + WM2cb: "\<^bold>\\<^bold>C\<^sub>\ (\<^bold>\(\<^sup>Aws c) \<^bold>\ (\<^bold>K\<^sub>b (\<^bold>\(\<^sup>Aws c))))\<^bold>\" + +text \Positive introspection principles are implied.\ +lemma WM2ab': "\<^bold>\\<^bold>C\<^sub>\ ((\<^sup>Aws a) \<^bold>\ \<^bold>K\<^sub>b (\<^sup>Aws a))\<^bold>\" + using WM2ab group_S5 unfolding Defs by metis +lemma WM2ac': "\<^bold>\\<^bold>C\<^sub>\ ((\<^sup>Aws a) \<^bold>\ \<^bold>K\<^sub>c (\<^sup>Aws a))\<^bold>\" + using WM2ac group_S5 unfolding Defs by metis +lemma WM2ba': "\<^bold>\\<^bold>C\<^sub>\ ((\<^sup>Aws b) \<^bold>\ \<^bold>K\<^sub>a (\<^sup>Aws b))\<^bold>\" + using WM2ba group_S5 unfolding Defs by metis +lemma WM2bc': "\<^bold>\\<^bold>C\<^sub>\ ((\<^sup>Aws b) \<^bold>\ \<^bold>K\<^sub>c (\<^sup>Aws b))\<^bold>\" + using WM2bc group_S5 unfolding Defs by metis +lemma WM2ca': "\<^bold>\\<^bold>C\<^sub>\ ((\<^sup>Aws c) \<^bold>\ \<^bold>K\<^sub>a (\<^sup>Aws c))\<^bold>\" + using WM2ca group_S5 unfolding Defs by metis +lemma WM2cb': "\<^bold>\\<^bold>C\<^sub>\ ((\<^sup>Aws c) \<^bold>\ \<^bold>K\<^sub>b (\<^sup>Aws c))\<^bold>\" + using WM2cb group_S5 unfolding Defs by metis + +text \Automated solutions of the Wise Men Puzzle.\ +theorem whitespot_c: "\<^bold>\\<^bold>[\<^bold>!\<^bold>\\<^bold>K\<^sub>a(\<^sup>Aws a)\<^bold>](\<^bold>[\<^bold>!\<^bold>\\<^bold>K\<^sub>b(\<^sup>Aws b)\<^bold>](\<^bold>K\<^sub>c (\<^sup>Aws c)))\<^bold>\" + using WM1 WM2ba WM2ca WM2cb unfolding Defs by (smt (verit)) + +text \For the following, alternative formulation a proof is found by sledgehammer, while the +reconstruction of this proof using trusted methods (often) fails; this hints at further opportunities to +improve the reasoning tools in Isabelle/HOL.\ +theorem whitespot_c': + "\<^bold>\\<^bold>[\<^bold>!\<^bold>\((\<^bold>K\<^sub>a (\<^sup>Aws a)) \<^bold>\ (\<^bold>K\<^sub>a (\<^bold>\\<^sup>Aws a)))\<^bold>](\<^bold>[\<^bold>!\<^bold>\((\<^bold>K\<^sub>b (\<^sup>Aws b)) \<^bold>\ (\<^bold>K\<^sub>b (\<^bold>\\<^sup>Aws b)))\<^bold>](\<^bold>K\<^sub>c (\<^sup>Aws c)))\<^bold>\" + using WM1 WM2ab WM2ac WM2ba WM2bc WM2ca WM2cb unfolding Defs + \ \sledgehammer by (smt (verit))\ + oops + +text \Consistency: nitpick reports a model.\ +lemma True nitpick [satisfy] oops +end + + + + + diff --git a/thys/PAL/ROOT b/thys/PAL/ROOT new file mode 100644 --- /dev/null +++ b/thys/PAL/ROOT @@ -0,0 +1,9 @@ +chapter AFP + +session "PAL" (AFP) = HOL + + options [timeout = 300] + theories + PAL + document_files + "root.bib" + "root.tex" diff --git a/thys/PAL/document/root.bib b/thys/PAL/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/PAL/document/root.bib @@ -0,0 +1,54 @@ +@techreport{R78, + Author = {Benzm{\"u}ller, Christoph and Reiche, Sebastian}, + Keywords = {own, Higher Order Logic, Hintikka Sets, Completeness Proofs}, + Title = {Modeling and Automating Public Announcement Logic with Relativized Common Knowledge as a Fragment of {HOL} in {LogiKEy}}, + Institution = {CoRR}, + Number = {arXiv:2111.01654}, + eprint = {2111.01654}, + Addendum = {\url{https://arxiv.org/abs/2111.01654}}, + Url = {https://arxiv.org/abs/2004.07506}, + Year = 2021 +} + +@InProceedings{C90, + Keywords = {own, Artificial Intelligence, Non-Classical Logics}, + Author = {Sebastian Reiche and Christoph Benzm{\"u}ller}, + Title = {Public Announcement Logic in {HOL}}, + Addendum = {Preprint: \url{https://arxiv.org/abs/2010.00810}}, + Url = {https://doi.org/10.1007/978-3-030-65840-3_14}, + Doi = {10.1007/978-3-030-65840-3_14}, + Booktitle = {Dynamic Logic. New Trends and Applications. DaLi 2020}, + editor = {Martins, Manuel A. and Sedl\'{a}r Igor}, + series = {Lecture Notes in Computer Science}, + volume = {12569}, + publisher = {Springer, Cham}, + year = 2020, + isbn = {978-3-030-65839-7}, +} + +@article{J41, + author = {Christoph Benzm{\"u}ller}, + title = {Universal (Meta-)Logical Reasoning: Recent + Successes}, + journal = {Science of Computer Programming}, + year = 2019, + volume = 172, + pages = {48-62}, + Addendum = {Preprint: + \url{http://doi.org/10.13140/RG.2.2.11039.61609/2}}, + doi = {10.1016/j.scico.2018.10.008}, +} + +@article{J44, + Author = {Christoph Benzm\"uller}, + Journal = {Data in Brief}, + Keywords = {own, Universal Logical Reasoning}, + Addendum = {Note: data publication}, + Title = {Universal (Meta-)Logical Reasoning: The Wise Men Puzzle ({Isabelle/HOL Dataset})}, + Year = 2019, + Volume = 24, + Pages = {1--5}, + Number = 103823, + OptAddendum = {Open Access}, + Doi = {10.1016/j.dib.2019.103823}, +} \ No newline at end of file diff --git a/thys/PAL/document/root.tex b/thys/PAL/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/PAL/document/root.tex @@ -0,0 +1,33 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% this should be the last package used +\usepackage{pdfsetup} +\usepackage{stmaryrd} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + + +\begin{document} + +\title{Automating Public Announcement Logic and the Wise Men Puzzle in +Isabelle/HOL} +\author{Christoph Benzm\"uller and Sebastian Reiche} +\maketitle + +\begin{abstract} +We present a shallow embedding of public announcement logic (PAL) with relativized general knowledge in HOL. We then use PAL to obtain an elegant encoding of the wise men puzzle, which we solve automatically using sledgehammer. +\end{abstract} + +\tableofcontents + + +% include generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,636 +1,640 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CSP_RefTK CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CoCon CofGroups Coinductive Coinductive_Languages Collections Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras CoSMed CoSMeDis Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Cubic_Quartic_Equations DFS_Framework DOM_Components DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 +Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker +PAL PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl +Real_Power Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Schutz_Spacetime Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szpilrajn +Szemeredi_Regularity TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weighted_Path_Order Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL diff --git a/thys/Real_Power/Log.thy b/thys/Real_Power/Log.thy new file mode 100644 --- /dev/null +++ b/thys/Real_Power/Log.thy @@ -0,0 +1,221 @@ +(* Title: RealPower/Log.thy + Authors: Jacques D. Fleuriot + University of Edinburgh, 2021 +*) + +section\Real Logarithms (Redefined)\ + +theory Log +imports RealPower +begin + +text\We can now directly define real logarithm of @{term x} to base @{term a}.\ + +definition + Log :: "[real,real] \ real" where + "Log a x = (THE y. a pow\<^sub>\ y = x)" + +lemma IVT_simple: + "\f (a::real) \ (y::real); y \ f b; a \ b; + \x. a \ x \ x \ b \ isCont f x\ + \ \x. f x = y" +by (frule IVT [of f]) auto + +lemma inj_on_powreal: + "0 < a \ a \ 1 \ inj_on (\x. a pow\<^sub>\ x) UNIV" +by (auto simp add: inj_on_def) + +lemma LIMSEQ_powreal_minus_nat: + "a > 1 \ (\n. a pow\<^sub>\ (-real n)) \ 0" +by (simp add: powreal_minus powreal_power_eq + LIMSEQ_inverse_realpow_zero) + +lemma LIMSEQ_less_Ex: + "\ X \ (x::real); x < y \ \ \n. X n < y" + by (meson LIMSEQ_le_const not_less) + +lemma powreal_IVT_upper_lemma: + assumes "a > (1::real)" and "x > 0" + shows "\n::nat. a pow\<^sub>\ (-real n) < x" +proof - + have "(\n. a pow\<^sub>\ - real n) \ 0" + by (simp add: LIMSEQ_powreal_minus_nat assms(1)) + then show ?thesis + using LIMSEQ_less_Ex assms(2) by blast +qed + +lemma powreal_IVT_lower_lemma: + assumes "a > (1::real)" + and "x > 0" + shows "\n::nat. x < a pow\<^sub>\ (real n)" +proof - + have invx0: "0 < inverse x" + by (simp add: assms(2)) + then have "\n. a pow\<^sub>\ - real n < inverse x" + using assms(1) powreal_IVT_upper_lemma by blast + then show ?thesis + using assms(1) + by (auto dest: inverse_less_imp_less + simp add: powreal_minus powreal_gt_zero ) +qed + +lemma powreal_surj: + assumes "a > 1" + and "x > 0" + shows "\y. a pow\<^sub>\ y = x" +proof - + obtain n where "a pow\<^sub>\ - real n < x" + using assms powreal_IVT_upper_lemma by blast + moreover obtain na where "x < a pow\<^sub>\ real na" + using assms powreal_IVT_lower_lemma by blast + moreover have "\x. - real n \ x \ x \ real na \ isCont ((pow\<^sub>\) a) x" + using assms(1) isCont_powreal_exponent_gt_one by blast + ultimately show ?thesis + using IVT_simple [of _ "-real n" _ "real na"] by force +qed + +lemma powreal_surj2: + "\ 0 < a; a < 1; x > 0 \ \ \y. a pow\<^sub>\ y = x" + using powreal_minus_base_ge_one powreal_surj real_inverse_gt_one_lemma + by blast + +lemma powreal_ex1_eq: + assumes "a > 0" + and "a \ 1" + and "x > 0" + shows "\! y. a pow\<^sub>\ y = x" +proof (cases "a < 1") + case True + then show ?thesis + using assms powreal_inject powreal_surj2 by blast +next + case False + then show ?thesis + using assms(2) assms(3) powreal_surj by auto +qed + +lemma powreal_Log_cancel [simp]: + "\ a > 0; a \ 1; x > 0 \ \ a pow\<^sub>\ (Log a x) = x" +by (auto intro: the1I2 [OF powreal_ex1_eq] simp add: Log_def) + +lemma Log_powreal_cancel [simp]: + "\ 0 < a; a \ 1 \ \ Log a (a pow\<^sub>\ y) = y" +by (metis powreal_ex1_eq powreal_gt_zero powreal_Log_cancel) + +lemma Log_mult: + "\ 0 < a; a \ 1; 0 < x; 0 < y \ + \ Log a (x * y) = Log a x + Log a y" + by (metis Log_powreal_cancel powreal_Log_cancel powreal_add) + +lemma Log_one [simp]: "\ 0 < a; a \ 1 \ \ Log a 1 = 0" +by (metis Log_powreal_cancel powreal_zero_eq_one) + +lemma Log_eq_one [simp]: "\ 0 < a; a \ 1 \ \ Log a a = 1" + using powreal_inject by fastforce + +lemma Log_inverse: + "\ a > 0; a \ 1; x > 0 \ \ Log a (inverse x) = - Log a x" +by (metis Log_powreal_cancel powreal_Log_cancel powreal_minus) + +lemma Log_divide: + "\ 0 < a; a \ 1; 0 < x; 0 < y \ + \ Log a (x/y) = Log a x - Log a y" + by (metis Log_inverse Log_mult divide_real_def + inverse_positive_iff_positive minus_real_def) + +lemma Log_less_cancel_iff [simp]: + assumes "1 < a" + and "0 < x" + and "0 < y" +shows "(Log a x < Log a y) = (x < y)" +proof + assume "Log a x < Log a y" + then show "x < y" using powreal_Log_cancel assms powreal_less_cancel_iff + by (metis less_irrefl real_inverse_bet_one_one_lemma + inverse_positive_iff_positive) +next + assume "x < y" + then show "Log a x < Log a y" + using assms(1) assms(2) powreal_less_cancel_iff by fastforce +qed + +lemma Log_inj: assumes "1 < b" shows "inj_on (Log b) {0 <..}" +proof (rule inj_onI, simp) + fix x y assume pos: "0 < x" "0 < y" and *: "Log b x = Log b y" + show "x = y" + proof (cases rule: linorder_cases) + assume "x < y" hence "Log b x < Log b y" + using Log_less_cancel_iff[OF `1 < b`] pos by simp + thus ?thesis using * by simp + next + assume "y < x" hence "Log b y < Log b x" + using Log_less_cancel_iff[OF `1 < b`] pos by simp + thus ?thesis using * by simp + qed simp +qed + +lemma Log_le_cancel_iff [simp]: + "\ 1 < a; 0 < x; 0 < y \ \ (Log a x \ Log a y) = (x \ y)" +by (simp add: linorder_not_less [symmetric]) + +lemma zero_less_Log_cancel_iff [simp]: + "1 < a \ 0 < x \ 0 < Log a x \ 1 < x" + using Log_less_cancel_iff[of a 1 x] by simp + +lemma zero_le_Log_cancel_iff[simp]: + "1 < a \ 0 < x \ 0 \ Log a x \ 1 \ x" + using Log_le_cancel_iff[of a 1 x] by simp + +lemma Log_less_zero_cancel_iff[simp]: + "1 < a \ 0 < x \ Log a x < 0 \ x < 1" + using Log_less_cancel_iff[of a x 1] by simp + +lemma Log_le_zero_cancel_iff[simp]: + "1 < a \ 0 < x \ Log a x \ 0 \ x \ 1" + using Log_le_cancel_iff[of a x 1] by simp + +lemma one_less_Log_cancel_iff[simp]: + "1 < a \ 0 < x \ 1 < Log a x \ a < x" + using Log_less_cancel_iff[of a a x] by simp + +lemma one_le_Log_cancel_iff[simp]: + "1 < a \ 0 < x \ 1 \ Log a x \ a \ x" + using Log_le_cancel_iff[of a a x] by simp + +lemma Log_less_one_cancel_iff[simp]: + "1 < a \ 0 < x \ Log a x < 1 \ x < a" + using Log_less_cancel_iff[of a x a] by simp + +lemma Log_le_one_cancel_iff[simp]: + "1 < a \ 0 < x \ Log a x \ 1 \ x \ a" + using Log_le_cancel_iff[of a x a] by simp + +lemma Log_powreal: + assumes "0 < x" + and "1 < b" + and "b \ 1" +shows "Log b (x pow\<^sub>\ y) = y * Log b x" +proof - + have "b pow\<^sub>\ (Log b x * y) = x pow\<^sub>\ y" + using assms powreal_mult [symmetric] by simp + moreover have "0 < x pow\<^sub>\ y" + by (simp add: assms(1) powreal_gt_zero) + ultimately have "b pow\<^sub>\ (y * Log b x) = b pow\<^sub>\ Log b (x pow\<^sub>\ y)" + using powreal_Log_cancel assms powreal_Log_cancel + by (simp add: mult.commute) + then show ?thesis + using assms(2) powreal_inject_exp1 by blast +qed + +lemma Log_nat_power: + assumes "0 < x" + and "1 < b" and "b \ 1" + shows " Log b (x ^ n) = real n * Log b x" +proof - + have "Log b (x pow\<^sub>\ real n) = real n * Log b x" + by (simp add: Log_powreal assms) + then show ?thesis + by (simp add: assms(1) powreal_power_eq) +qed + +end \ No newline at end of file diff --git a/thys/Real_Power/ROOT b/thys/Real_Power/ROOT new file mode 100644 --- /dev/null +++ b/thys/Real_Power/ROOT @@ -0,0 +1,8 @@ +chapter AFP + +session "Real_Power" (AFP) = HOL + + options [timeout = 300] + theories + Log + document_files + "root.tex" diff --git a/thys/Real_Power/RatPower.thy b/thys/Real_Power/RatPower.thy new file mode 100644 --- /dev/null +++ b/thys/Real_Power/RatPower.thy @@ -0,0 +1,749 @@ +(* Title: RealPower/RatPower.thy + Authors: Jacques D. Fleuriot + University of Edinburgh, 2021 +*) + +theory RatPower +imports HOL.NthRoot +begin + +section \Rational Exponents\ + +text\A few lemmas about nth-root.\ + +lemma real_root_mult_exp_cancel: + "\ 0 < x; 0 < m; 0 < n \ + \ root (m * n) (x ^ (k * n)) = root m (x ^ k)" + by (simp add: power_mult real_root_pos_unique) + +lemma real_root_mult_exp_cancel1: + "\ 0 < x; 0 < n \ \ root n (x ^ (k * n)) = x ^ k" +by (auto dest: real_root_mult_exp_cancel [of _ 1]) + +lemma real_root_mult_exp_cancel2: + "\ 0 < x; 0 < m; 0 < n \ + \ root (n * m) (x ^ (n * k)) = root m (x ^ k)" +by (simp add: mult.commute real_root_mult_exp_cancel) + +lemma real_root_mult_exp_cancel3: + "\ 0 < x; 0 < n \ \ root n (x ^ (n * k)) = x ^ k" +by (auto dest: real_root_mult_exp_cancel2 [of _ 1]) + +text\Definition of rational exponents,\ + +definition + powrat :: "[real,rat] => real" (infixr "pow\<^sub>\" 80) where + "x pow\<^sub>\ r = (if r > 0 + then root (nat (snd(quotient_of r))) + (x ^ (nat (fst(quotient_of r)))) + else root (nat (snd(quotient_of r))) + (1/x ^ (nat (- fst(quotient_of r)))))" + +(* Why isn't this a default simp rule? *) +declare quotient_of_denom_pos' [simp] + +lemma powrat_one_eq_one [simp]: "1 pow\<^sub>\ a = 1" + by (simp add: powrat_def) + +lemma powrat_zero_eq_one [simp]: "x pow\<^sub>\ 0 = 1" +by (simp add: powrat_def) + +lemma powrat_one [simp]: "x pow\<^sub>\ 1 = x" +by (simp add: powrat_def) + +lemma powrat_mult_base: + "(x * y) pow\<^sub>\ r = (x pow\<^sub>\ r) * (y pow\<^sub>\ r)" +proof (cases r) + case (Fract a b) + then show ?thesis + using powrat_def quotient_of_Fract real_root_mult [symmetric] + power_mult_distrib by fastforce +qed + +lemma powrat_divide: + "(x / y) pow\<^sub>\ r = (x pow\<^sub>\ r)/(y pow\<^sub>\ r)" +proof (cases r) + case (Fract a b) + then show ?thesis + using powrat_def quotient_of_Fract real_root_divide [symmetric] + power_divide by fastforce +qed + +lemma powrat_zero_base [simp]: + assumes "r \ 0" shows "0 pow\<^sub>\ r = 0" +proof (cases r) + case (Fract a b) + then show ?thesis + proof (cases "a > 0") + case True + then show ?thesis + using Fract powrat_def quotient_of_Fract zero_less_Fract_iff + by simp + next + case False + then have "a \ 0" + using Fract(1) assms rat_number_collapse(1) by blast + then + show ?thesis + using Fract powrat_def quotient_of_Fract zero_less_Fract_iff + by auto + qed +qed + + +(* That's the one we want *) +lemma powrat_inverse: + "(inverse y) pow\<^sub>\ r = inverse(y pow\<^sub>\ r)" +proof (cases "r=0") + case True + then show ?thesis by simp +next + case False + then show ?thesis + by (simp add: inverse_eq_divide powrat_divide) +qed + +lemma powrat_minus: + "x pow\<^sub>\ (-r) = inverse (x pow\<^sub>\ r)" +proof (cases r) + case (Fract a b) + then show ?thesis + by (auto simp add: powrat_def divide_inverse real_root_inverse + quotient_of_Fract zero_less_Fract_iff) +qed + +lemma powrat_gt_zero: + assumes "x > 0" shows "x pow\<^sub>\ r > 0" +proof (cases r) + case (Fract a b) + then show ?thesis + by (simp add: assms powrat_def) +qed + +lemma powrat_not_zero: + assumes "x \ 0" shows "x pow\<^sub>\ r \ 0" +proof (cases r) + case (Fract a b) + then show ?thesis + by (simp add: assms powrat_def) +qed + +(* Not in GCD.thy *) +lemma gcd_add_mult_commute: "gcd (m::'a::semiring_gcd) (n + k * m) = gcd m n" + by (metis add.commute gcd_add_mult) + +lemma coprime_add_mult_iff1 [simp]: + "coprime (n + k * m) (m::'a::semiring_gcd) = coprime n m" + by (simp add: coprime_iff_gcd_eq_1 gcd.commute gcd_add_mult_commute) + +lemma coprime_add_mult_iff2 [simp]: + "coprime (k * m + n) (m::'a::semiring_gcd) = coprime n m" + by (simp add: add.commute) + +(* Not proved before?? *) +lemma gcd_mult_div_cancel_left1 [simp]: + "gcd a b * (a div gcd a b) = (a::'a::semiring_gcd)" + by simp + +lemma gcd_mult_div_cancel_left2 [simp]: + "gcd b a * (a div gcd b a) = (a::'a::semiring_gcd)" + by simp + +lemma gcd_mult_div_cancel_right1 [simp]: + "(a div gcd a b) * gcd a b = (a::'a::semiring_gcd)" + by simp + +lemma gcd_mult_div_cancel_right2 [simp]: + "(a div gcd b a) * gcd b a = (a::'a::semiring_gcd)" + by simp +(* END: Not in GCD.thy *) + +lemma real_root_normalize_cancel: + assumes "0 < x" and "a \ 0" and "b > 0" + shows "root (nat(snd(Rat.normalize(a,b)))) + (x ^ nat(fst(Rat.normalize(a,b)))) = + root (nat b) (x ^ (nat a))" +proof - + have "root (nat (b div gcd a b)) (x ^ nat (a div gcd a b)) = + root (nat b) (x ^ nat a)" + proof (cases "coprime a b") + case True + then show ?thesis by simp + next + case False + have "0 < gcd a b" + using assms(2) gcd_pos_int by blast + then have "nat (gcd a b) > 0" + by linarith + moreover have "nat (b div gcd a b) > 0" + using nonneg1_imp_zdiv_pos_iff assms(3) by auto + moreover + have "root (nat b) (x ^ nat a) = + root (nat (gcd a b * (b div gcd a b))) + (x ^ nat (gcd a b * (a div gcd a b)))" + by simp + ultimately show ?thesis + using assms(1) gcd_ge_0_int nat_mult_distrib real_root_mult_exp_cancel2 + by presburger + qed + then show ?thesis + by (metis assms(3) fst_conv normalize_def snd_conv) + qed + +lemma powrat_add_pos: + assumes "0 < x" and "0 < r" and "0 < s" + shows "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" +proof (cases r) + case (Fract a b) + assume b0: "b > 0" and rf: "r = Fract a b" + then have a0: "a > 0" + using Fract(1) assms(2) zero_less_Fract_iff by blast + then show ?thesis + proof (cases s) + case (Fract c d) + assume d0: "d > 0" and sf: "s = Fract c d" + then have c0: "c > 0" + using assms(3) zero_less_Fract_iff by blast + then have bd0: "b * d > 0" + using b0 zero_less_mult_iff Fract(2) by blast + then show ?thesis + proof (cases "a * d > 0 \ c * b > 0") + case True + assume abcd: "a * d > 0 \ c * b > 0" + then have adcb0: "a * d + c * b > 0" by simp + have "x ^ nat (a * d + c * b) = + ((x ^ (nat a)) ^ nat d) * (x ^ (nat c)) ^ nat b" + using abcd nat_mult_distrib nat_add_distrib + zero_less_Fract_iff Fract(3) a0 c0 + by (simp add: power_mult power_add) + then have "root (nat b) + (root (nat d) + (x ^ nat (a * d + c * b))) = + root (nat b) + (root (nat d) + (((x ^ (nat a)) ^ nat d) * (x ^ (nat c)) ^ nat b))" + by simp + also have "... = root (nat b) + (root (nat d) ((x ^ (nat a)) ^ nat d) * + root (nat d) ((x ^ (nat c)) ^ nat b))" + by (simp add: Fract(2) real_root_mult) + also have "... = root (nat b) + (root (nat d) ((x ^ (nat a)) ^ nat d)) * + root (nat b) + (root (nat d) ((x ^ (nat c)) ^ nat b))" + using real_root_mult by blast + also have "... = root (nat b) (x ^ (nat a)) * + root (nat b) + (root (nat d) ((x ^ (nat c)) ^ nat b))" + using real_root_power_cancel Fract(2) zero_less_nat_eq assms(1) + less_imp_le by simp + also have "... = + root (nat b) (x ^ nat a) * root (nat d) (x ^ nat c)" + using real_root_power [of "nat d"] real_root_power_cancel + real_root_pos_pos_le zero_less_nat_eq + Fract(2) zero_less_nat_eq b0 assms(1) by auto + finally have "root (nat b) (root (nat d) (x ^ nat (a * d + c * b))) = + root (nat b) (x ^ nat a) * root (nat d) (x ^ nat c)" + by assumption + then show ?thesis using a0 b0 c0 d0 bd0 abcd adcb0 assms(1) sf rf + by (auto simp add: powrat_def quotient_of_Fract + real_root_mult_exp nat_mult_distrib + zero_less_Fract_iff real_root_normalize_cancel) + next + case False + then show ?thesis + by (simp add: a0 b0 c0 d0) + qed + qed +qed + +lemma powrat_add_neg: + assumes "0 < x" and "r < 0" and "s < 0" + shows "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" +proof - + have "x pow\<^sub>\ (- r + - s) = x pow\<^sub>\ - r * x pow\<^sub>\ - s" + using assms powrat_add_pos neg_0_less_iff_less by blast + then show ?thesis + by (metis inverse_eq_imp_eq inverse_mult_distrib + minus_add_distrib powrat_minus) +qed + +lemma powrat_add_neg_pos: + assumes pos_x: "0 < x" and + neg_r: "r < 0" and + pos_s: "0 < s" + shows "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" +proof (cases "r + s > 0") + assume exp_pos: "r + s > 0" + have "-r > 0" using neg_r by simp + then have "x pow\<^sub>\ (r + s) * (x pow\<^sub>\ -r) = (x pow\<^sub>\ s)" + using exp_pos pos_x powrat_add_pos by fastforce + then have "x pow\<^sub>\ (r + s) * inverse (x pow\<^sub>\ r) = (x pow\<^sub>\ s)" + by (simp add: powrat_minus) + then show "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" + by (metis Groups.mult_ac(3) assms(1) mult.right_neutral + order_less_irrefl powrat_gt_zero right_inverse) +next + assume exp_pos: "\ r + s > 0" + then have "r + s = 0 \ r + s < 0" using neq_iff by blast + then show "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" + proof + assume exp_zero: "r + s = 0" + then have "x pow\<^sub>\ (r + s) = 1" by simp + also have "... = (x pow\<^sub>\ r) * inverse (x pow\<^sub>\ r)" + using pos_x powrat_not_zero by simp + also have "... = (x pow\<^sub>\ r) * (x pow\<^sub>\ - r)" + by (simp add: powrat_minus) + finally show "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" + using exp_zero minus_unique by blast + next + assume "r + s < 0" + have "-s < 0" using pos_s by simp + then have "x pow\<^sub>\ (r + s) * (x pow\<^sub>\ -s) = (x pow\<^sub>\ r)" + using \r + s < 0\ pos_x powrat_add_neg by fastforce + then have "x pow\<^sub>\ (r + s) * inverse (x pow\<^sub>\ s) = (x pow\<^sub>\ r)" + by (simp add: powrat_minus) + then show "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" + by (metis divide_eq_eq divide_real_def + less_irrefl pos_x powrat_not_zero) + qed +qed + +lemma powrat_add_pos_neg: + "\ 0 < x; 0 < r; s < 0 \ + \ x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" +by (metis add.commute mult.commute powrat_add_neg_pos) + +lemma powrat_add: + assumes "0 < x" + shows "x pow\<^sub>\ (r + s) = (x pow\<^sub>\ r) * (x pow\<^sub>\ s)" + proof (cases "(r > 0 \ r \ 0) \ (s > 0 \ s \ 0)") + case True + then show ?thesis using assms + by (auto dest: powrat_add_pos powrat_add_neg powrat_add_neg_pos + powrat_add_pos_neg simp add: le_less) +next + case False + then show ?thesis by auto +qed + +lemma powrat_diff: + "0 < x \ x pow\<^sub>\ (a - b) = x pow\<^sub>\ a / x pow\<^sub>\ b" +by (metis add_uminus_conv_diff divide_inverse powrat_add powrat_minus) + +lemma powrat_mult_pos: + assumes "0 < x" and "0 < r" and "0 < s" + shows "x pow\<^sub>\ (r * s) = (x pow\<^sub>\ r) pow\<^sub>\ s" +proof (cases r) + case (Fract a b) + assume b0: "b > 0" and rf: "r = Fract a b" and coab: "coprime a b" + have a0: "a > 0" + using assms(2) b0 rf zero_less_Fract_iff by blast + then show ?thesis + proof (cases s) + case (Fract c d) + assume d0: "d > 0" and sf: "s = Fract c d" and coad: "coprime c d" + then have c0: "c > 0" + using assms(3) zero_less_Fract_iff by blast + then have "b * d > 0" + using b0 d0 by simp + then show ?thesis using a0 c0 b0 d0 rf sf coab coad assms + by (auto intro: mult_pos_pos simp add: powrat_def quotient_of_Fract + zero_less_Fract_iff real_root_normalize_cancel real_root_power + real_root_mult_exp [symmetric] nat_mult_distrib power_mult + mult.commute) + qed +qed + +lemma powrat_mult_neg: + assumes "0 < x" "r < 0" and "s < 0" + shows "x pow\<^sub>\ (r * s) = (x pow\<^sub>\ r) pow\<^sub>\ s" +proof - + have " x pow\<^sub>\ (- r * - s) = (x pow\<^sub>\ - r) pow\<^sub>\ - s" + using powrat_mult_pos assms neg_0_less_iff_less by blast + then show ?thesis + by (simp add: powrat_inverse powrat_minus) +qed + + +lemma powrat_mult_neg_pos: + assumes "0 < x" and "r < 0" and "0 < s" + shows "x pow\<^sub>\ (r * s) = (x pow\<^sub>\ r) pow\<^sub>\ s" +proof - + have "x pow\<^sub>\ (- r * s) = (x pow\<^sub>\ - r) pow\<^sub>\ s" + using powrat_mult_pos assms neg_0_less_iff_less by blast + then show ?thesis + by (simp add: powrat_inverse powrat_minus) +qed + +lemma powrat_mult_pos_neg: + assumes "0 < x" and "0 < r" and "s < 0" + shows "x pow\<^sub>\ (r * s) = (x pow\<^sub>\ r) pow\<^sub>\ s" +proof - + have "x pow\<^sub>\ (r * - s) = (x pow\<^sub>\ r) pow\<^sub>\ - s" + using powrat_mult_pos assms neg_0_less_iff_less by blast + then show ?thesis + by (simp add: powrat_minus) +qed + +lemma powrat_mult: + assumes "0 < x" shows "x pow\<^sub>\ (r * s) = (x pow\<^sub>\ r) pow\<^sub>\ s" +proof - + {fix q::rat + assume "q = 0" then have "x pow\<^sub>\ (q * s) = (x pow\<^sub>\ q) pow\<^sub>\ s" + by simp + } + then show ?thesis + by (metis assms linorder_neqE_linordered_idom mult_zero_right + powrat_mult_neg powrat_mult_neg_pos powrat_mult_pos + powrat_mult_pos_neg powrat_zero_eq_one) +qed + +lemma powrat_less_mono: + assumes "r < s" and "1 < x" + shows "x pow\<^sub>\ r < x pow\<^sub>\ s" + proof (cases r) + case (Fract a b) + assume r_assms: "r = Fract a b" "0 < b" "coprime a b" + then show ?thesis + proof (cases s) + case (Fract c d) + assume s_assms: "s = Fract c d" "0 < d" "coprime c d" + have adcb: "a * d < c * b" + using assms r_assms s_assms by auto + have b_ba: "0 < nat (b * d)" + by (simp add: r_assms(2) s_assms(2)) + have root0: "0 \ root (nat d) (x ^ nat c)" + "0 \ root (nat d) (1 / x ^ nat (- c))" + using assms(2) real_root_pos_pos_le by auto + then show ?thesis + proof (auto simp add: powrat_def quotient_of_Fract zero_less_Fract_iff + s_assms r_assms) + assume ac0: "0 < a" "0 < c" + then have "(x ^ nat a) ^ nat d < (x ^ nat c) ^ nat b" + using adcb assms(2) r_assms(2) less_imp_le mult_pos_pos + nat_mono_iff nat_mult_distrib power_mult + power_strict_increasing + by metis + then have "root (nat b) (x ^ nat a) ^ nat (b * d) < + root (nat d) (x ^ nat c) ^ nat (b * d)" + using assms r_assms s_assms ac0 real_root_power + [symmetric] nat_mult_distrib + by (auto simp add: power_mult) + then show "root (nat b) (x ^ nat a) < root (nat d) (x ^ nat c)" + using power_less_imp_less_base root0(1) by blast + next + assume "0 < a" "\ 0 < c" + then show "root (nat b) (x ^ nat a) < + root (nat d) (1 / x ^ nat (- c))" + using assms(1) less_trans r_assms(1) r_assms(2) s_assms(1) + s_assms(2) zero_less_Fract_iff by blast + next + assume ac0: "\ 0 < a" "0 < c" + then have "a = 0 \ a < 0" by auto + then show "root (nat b) (1 / x ^ nat (- a)) < + root (nat d) (x ^ nat c)" + proof + assume "a = 0" then show ?thesis + by (simp add: ac0(2) assms(2) r_assms(2) s_assms(2)) + next + assume a0: "a < 0" + have "(1 / x ^ nat (- a)) ^ nat d < 1" + using a0 s_assms(2) assms(2) adcb power_mult [symmetric] + power_one_over nat_mult_distrib [symmetric] + by (metis (no_types) eq_divide_eq_1 inverse_eq_divide inverse_le_1_iff + le_less neg_0_less_iff_less one_less_power power_one_over + zero_less_nat_eq) + moreover have "1 < (x ^ nat c) ^ nat b" + by (simp add: ac0(2) assms(2) r_assms(2)) + ultimately have "(1 / x ^ nat (- a)) ^ nat d < (x ^ nat c) ^ nat b" + by linarith + then have "root (nat b) (1 / x ^ nat (- a)) ^ nat (b * d) + < root (nat d) (x ^ nat c) ^ nat (b * d)" + using assms r_assms s_assms ac0 real_root_power [symmetric] nat_mult_distrib + by (auto simp add: power_mult) + then show ?thesis + using power_less_imp_less_base root0(1) by blast + qed + next + assume ac0: "\ 0 < a" "\ 0 < c" + then show "root (nat b) (1 / x ^ nat (- a)) < + root (nat d) (1 / x ^ nat (- c))" + proof (cases "a = 0 \ c = 0") + assume "a = 0 \ c = 0" then show ?thesis + using assms r_assms s_assms adcb ac0 + by (auto simp add: not_less le_less) + next + assume "\ (a = 0 \ c = 0)" + then have ac00: "a < 0" "c < 0" using ac0 by auto + then have "1 / x ^ nat (- (a * d)) < + 1 / x ^ nat (- (c * b))" + using ac00 r_assms s_assms assms + by (simp add: divide_inverse mult_pos_neg2) + then have "(1 / x ^ nat (- a)) ^ nat d < + (1 / x ^ nat (- c)) ^ nat b" + using s_assms r_assms ac00 + by (auto simp add: power_mult [symmetric] + power_one_over nat_mult_distrib [symmetric]) + then have "root (nat b) (1 / x ^ nat (- a)) ^ nat (b * d) + < root (nat d) (1 / x ^ nat (- c)) ^ nat (b * d)" + using assms r_assms s_assms ac0 real_root_power [symmetric] + nat_mult_distrib + by (auto simp add: power_mult) + then show "root (nat b) (1 / x ^ nat (- a)) < + root (nat d) (1 / x ^ nat (- c))" + using power_less_imp_less_base root0(2) by blast + qed + qed + qed +qed + +lemma power_le_imp_le_base2: + "\ (a::'a::linordered_semidom) ^ n \ b ^ n; 0 \ b; 0 < n \ + \ a \ b" +by (auto intro: power_le_imp_le_base [of _ "n - 1"]) + +lemma powrat_le_mono: + assumes "r \ s" and "1 \ x" + shows "x pow\<^sub>\ r \ x pow\<^sub>\ s" + by (metis (full_types) assms le_less powrat_less_mono powrat_one_eq_one) + +lemma powrat_less_cancel: + "\ x pow\<^sub>\ r < x pow\<^sub>\ s; 1 < x \ \ r < s" + by (metis not_less_iff_gr_or_eq powrat_less_mono) + +(* Monotonically increasing *) +lemma powrat_less_cancel_iff [simp]: + "1 < x \ (x pow\<^sub>\ r < x pow\<^sub>\ s) = (r < s)" +by (blast intro: powrat_less_cancel powrat_less_mono) + +lemma powrat_le_cancel_iff [simp]: + "1 < x \ (x pow\<^sub>\ r \ x pow\<^sub>\ s) = (r \ s)" +by (simp add: linorder_not_less [symmetric]) + +(* Next 2 theorems should be in Power.thy *) +lemma power_inject_exp_less_one [simp]: + "\0 < a; (a::'a::{linordered_field}) < 1 \ + \ a ^ m = a ^ n \ m = n" +by (metis less_irrefl nat_neq_iff power_strict_decreasing) + +lemma power_inject_exp_strong [simp]: + "\0 < a; (a::'a::{linordered_field}) \ 1 \ + \ a ^ m = a ^ n \ m = n" +by (case_tac "a < 1") (auto simp add: not_less) + +(* Not proved elsewhere? *) +lemma nat_eq_cancel: "0 < a \ 0 < b \ (nat a = nat b) = (a = b)" +by auto + +lemma powrat_inject_exp [simp]: + "1 < x \ (x pow\<^sub>\ r = x pow\<^sub>\ s) = (s = r)" + by (metis neq_iff powrat_less_cancel_iff) + +lemma powrat_inject_exp_less_one [simp]: + assumes "0 < x" and "x < 1" + shows "(x pow\<^sub>\ r = x pow\<^sub>\ s) = (s = r)" +proof - + have "1 < inverse x" + using assms one_less_inverse by blast + then show ?thesis + using powrat_inject_exp powrat_inverse by fastforce +qed + +lemma powrat_inject_exp_strong [simp]: + "\ 0 < x; x \ 1 \ \ (x pow\<^sub>\ r = x pow\<^sub>\ s) = (s = r)" + using powrat_inject_exp_less_one by fastforce + +(* Monotonically decreasing *) +lemma powrat_less_1_cancel_iff [simp]: + assumes x0: "0 < x" and x1: "x < 1" + shows "(x pow\<^sub>\ r < x pow\<^sub>\ s) = (s < r)" +proof + assume xrs: "x pow\<^sub>\ r < x pow\<^sub>\ s" + have invx: "1 < 1/x" using assms by simp + have "r < s \ r \ s" using leI by blast + then show "s < r" + proof + assume "r < s" + then have " inverse x pow\<^sub>\ r < inverse x pow\<^sub>\ s" using invx + by (simp add: inverse_eq_divide) + then have " x pow\<^sub>\ s < x pow\<^sub>\ r" + by (simp add: powrat_gt_zero powrat_inverse x0) + then show ?thesis using xrs by linarith + next + assume "r \ s" + then show ?thesis + using less_eq_rat_def xrs by blast + qed +next + assume sr: "s < r" + have invx: "1 < 1/x" using assms by simp + then have " inverse x pow\<^sub>\ s < inverse x pow\<^sub>\ r" using invx sr + by (simp add: inverse_eq_divide) + then show "x pow\<^sub>\ r < x pow\<^sub>\ s" + by (simp add: powrat_gt_zero powrat_inverse x0) +qed + +lemma powrat_le_1_cancel_iff [simp]: + "\0 < x; x < 1\ \ (x pow\<^sub>\ r \ x pow\<^sub>\ s) = (s \ r)" +by (auto simp add: le_less) + +lemma powrat_ge_one: "x \ 1 \ r \ 0 \ x pow\<^sub>\ r \ 1" +by (metis powrat_le_mono powrat_zero_eq_one) + +lemma isCont_powrat: + assumes "0 < x" shows "isCont (\x. x pow\<^sub>\ r) x" +proof (cases r) + case (Fract a b) + assume fract_assms: "r = Fract a b" "0 < b" "coprime a b" + then show ?thesis + proof (cases "0 < a") + case True + then show ?thesis + using fract_assms isCont_o2 [OF isCont_power [OF continuous_ident]] + by (auto intro: isCont_real_root simp add: powrat_def zero_less_Fract_iff) + next + case False + then show ?thesis + using fract_assms assms isCont_real_root real_root_gt_zero + continuous_at_within_inverse [intro!] + by (auto intro!: isCont_o2 [OF isCont_power [OF continuous_ident]] + simp add: powrat_def zero_less_Fract_iff divide_inverse + real_root_inverse) + qed +qed + +lemma LIMSEQ_powrat_base: + "\ X \ a; a > 0 \ \ (\n. (X n) pow\<^sub>\ q) \ a pow\<^sub>\ q" +by (metis isCont_tendsto_compose [where g="\x. x pow\<^sub>\ q"] isCont_powrat) + +lemma powrat_inverse_of_nat_ge_one [simp]: + "a \ 1 \ a pow\<^sub>\ (inverse (of_nat n)) \ 1" + by (simp add: powrat_ge_one) + +lemma powrat_inverse_of_nat_le_self [simp]: + assumes "1 \ a" shows "a pow\<^sub>\ inverse (rat_of_nat n) \ a" +proof - + have "inverse (rat_of_nat n) \ 1" + by (auto simp add: inverse_le_1_iff) + also have "a pow\<^sub>\ 1 \ a" by simp + ultimately show ?thesis + using assms powrat_le_mono by fastforce +qed + +(* This lemma used to be in Limits.thy *) +lemma BseqI2': "\n\N. norm (X n) \ K \ Bseq X" + using BfunI eventually_sequentially by blast + +lemma Bseq_powrat_inverse_of_nat_ge_one: + "a \ 1 \ Bseq (\n. a pow\<^sub>\ (inverse (of_nat n)))" +by (auto intro: BseqI2' [of 1 _ a] simp add: less_imp_le powrat_gt_zero) + +lemma decseq_powrat_inverse_of_nat_ge_one: + "a \ 1 \ decseq (\n. a pow\<^sub>\ (inverse (of_nat (Suc n))))" +unfolding decseq_def by (auto intro: powrat_le_mono) + +lemma convergent_powrat_inverse_Suc_of_nat_ge_one: + assumes "a \ 1" + shows "convergent (\n. a pow\<^sub>\ (inverse (of_nat (Suc n))))" +proof - + have "Bseq (\n. a pow\<^sub>\ inverse (rat_of_nat n))" + using Bseq_powrat_inverse_of_nat_ge_one assms by blast + then have "Bseq (\n. a pow\<^sub>\ inverse (rat_of_nat (Suc n)))" + using Bseq_ignore_initial_segment [of _ 1] by fastforce + also have "monoseq (\n. a pow\<^sub>\ inverse (rat_of_nat (Suc n)))" + using assms decseq_imp_monoseq decseq_powrat_inverse_of_nat_ge_one by blast + ultimately show ?thesis + using Bseq_monoseq_convergent by blast +qed + +lemma convergent_powrat_inverse_of_nat_ge_one: + assumes "a \ 1" shows "convergent (\n. a pow\<^sub>\ (inverse (of_nat n)))" +proof - + have "convergent (\n. a pow\<^sub>\ inverse (rat_of_nat (Suc n)))" + using convergent_powrat_inverse_Suc_of_nat_ge_one assms by blast + then obtain L where "(\n. a pow\<^sub>\ inverse (1 + rat_of_nat n)) \ L" + using convergent_def by auto + then have "(\n. a pow\<^sub>\ inverse (rat_of_nat (n + 1))) \ L" + by simp + then have "(\n. a pow\<^sub>\ inverse (rat_of_nat n)) \ L" + by (rule LIMSEQ_offset [of _ 1]) + then show ?thesis using convergent_def by auto +qed + +lemma LIMSEQ_powrat_inverse_of_nat_ge_one: + assumes "a \ 1" shows "(\n. a pow\<^sub>\ (inverse (of_nat n))) \ 1" +proof - + have "convergent(\n. a pow\<^sub>\ inverse (rat_of_nat n))" + using convergent_powrat_inverse_of_nat_ge_one assms by blast + then have "\L. L \ 0 \ (\n. a pow\<^sub>\ (inverse (of_nat n))) \ L" + using assms convergent_def powrat_inverse_of_nat_ge_one + LIMSEQ_le_const not_one_le_zero + by metis + then obtain L + where l0: "L \ 0" + and liml: "(\n. a pow\<^sub>\ (inverse (of_nat n))) \ L" + by blast + then have "(\n. a pow\<^sub>\ (inverse (of_nat n)) * + a pow\<^sub>\ (inverse (of_nat n))) \ L * L" + by (simp add: tendsto_mult) + then have "(\n. a pow\<^sub>\ (2 * inverse (of_nat n))) \ L * L" + using powrat_add [symmetric] assms by simp + also have "(2::nat) > 0" by simp + ultimately + have "(\n. a pow\<^sub>\ (2 * inverse (of_nat (n * 2)))) \ L * L" + using LIMSEQ_linear [of _ "L * L" 2] by blast + then have "(\n. a pow\<^sub>\ (inverse (of_nat n))) \ L * L" + by simp + then show ?thesis using liml using LIMSEQ_unique l0 + by fastforce +qed + +lemma LIMSEQ_powrat_inverse_of_nat_pos_less_one: + assumes a0: "0 < a" and a1: "a < 1" + shows "(\n. a pow\<^sub>\ (inverse (of_nat n))) \ 1" +proof - + have "inverse a > 1" using a0 a1 + using one_less_inverse by blast + then have "(\n. inverse a pow\<^sub>\ inverse (rat_of_nat n)) \ 1" + using LIMSEQ_powrat_inverse_of_nat_ge_one by simp + then have "(\n. inverse (a pow\<^sub>\ inverse (rat_of_nat n))) \ 1" + using powrat_inverse by simp + then have "(\x. 1 / inverse (a pow\<^sub>\ inverse (rat_of_nat x))) \ 1/1" + by (auto intro: tendsto_divide simp only:) + then show ?thesis by (auto simp add: divide_inverse) +qed + +lemma LIMSEQ_powrat_inverse_of_nat: + "a > 0 \ (\n. a pow\<^sub>\ (inverse (of_nat n))) \ 1" + by (metis LIMSEQ_powrat_inverse_of_nat_ge_one + LIMSEQ_powrat_inverse_of_nat_pos_less_one leI) + + +lemma real_root_eq_powrat_inverse: + assumes "n > 0" shows "root n x = x pow\<^sub>\ (inverse (of_nat n))" +proof (cases n) + case 0 + then show ?thesis + using assms by blast +next + case (Suc m) + assume nSuc: "n = Suc m" + then have "root (Suc m) x = root (nat (1 + int m)) x" + by (metis nat_int of_nat_Suc) + then show ?thesis + by (auto simp add: nSuc powrat_def of_nat_rat + zero_less_Fract_iff quotient_of_Fract) +qed + +lemma powrat_power_eq: + "0 < a \ a pow\<^sub>\ rat_of_nat n = a ^ n" +proof (induction n) +case 0 + then show ?case by simp +next + case (Suc n) + then show ?case using powrat_add by simp +qed + +end \ No newline at end of file diff --git a/thys/Real_Power/RealPower.thy b/thys/Real_Power/RealPower.thy new file mode 100644 --- /dev/null +++ b/thys/Real_Power/RealPower.thy @@ -0,0 +1,1507 @@ +(* Title: RealPower/RealPower.thy + Authors: Jacques D. Fleuriot + University of Edinburgh, 2021 +*) + +section \Real Exponents via Limits\ + +theory RealPower +imports RatPower +begin + +instance rat :: ab_group_add +by intro_classes + +instance rat :: field +by intro_classes + +instance rat :: comm_ring_1 +by intro_classes + +instantiation rat :: dist +begin +definition rat_dist_def: + "dist x y = of_rat (abs(y - x))" + +instance .. +end + +instantiation rat :: dist_norm +begin +definition rat_norm_def: + "norm (q::rat) = of_rat \q\" + +instance +by intro_classes (simp add: rat_dist_def rat_norm_def) +end + +instantiation rat :: metric_space +begin + +definition uniformity_rat_def [code del]: + "(uniformity :: (rat \ rat) filter) = + (INF e\{0 <..}. principal {((x::rat), y). dist x y < e})" + +definition open_rat_def [code del]: + "open (U :: rat set) \ + (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" + +instance +proof + show "uniformity = + (INF e\{0<..}. principal {(x::rat, y). dist x y < e})" + using uniformity_rat_def by blast +next + fix U + show "open U = + (\x\U. \\<^sub>F (x', y::rat) in uniformity. x' = x \ y \ U)" + using open_rat_def by blast +next + fix x y show "(dist x y = 0) = (x = (y::rat))" + by (force simp add: rat_dist_def) +next + fix x y z show "dist x y \ dist x z + dist y (z::rat)" + by (force simp add: rat_dist_def of_rat_add [symmetric] of_rat_less_eq) +qed + +end + +instance rat :: topological_space +by intro_classes + +lemma LIMSEQ_squeeze: + assumes abc: "\n. a n \ b n \ b n \ c n" + and alim: "a \ (L::real)" + and clim: "c \ L" shows "b \ L" +proof - + {fix r + assume r0: "(r::real) > 0" + then have "\no. \n\no. \a n - L\ < r" + by (metis LIMSEQ_def alim dist_real_def) + then obtain p where 1: "\n\p. \a n - L\ < r" by blast + have "\no. \n\no. \c n - L\ < r" + by (metis LIMSEQ_def r0 clim dist_real_def) + then obtain k where 2: "\n\k. \c n - L\ < r" by blast + have "\m. \n\m. \b n - L\ < r" + proof - + {fix n + assume n_max: "max p k \ n" + then have a_n: "\a n - L\ < r" using 1 by simp + have c_n: "\c n - L\ < r" using n_max 2 by simp + have "a n \ b n \ b n \ c n" using abc by simp + then have "\b n - L\ < r" + using a_n c_n by linarith + } + then have "\n\max p k. \b n - L\ < r" + by blast + then show ?thesis by blast + qed + } + then have "\r>0. \m. \n\m. \b n - L\ < r" + by blast + then show ?thesis + by (simp add: dist_real_def lim_sequentially) +qed + +primrec incratseq :: "real \ nat \ rat" where + "incratseq x 0 = (\q. x - 1 < of_rat q \ of_rat q < x)" +| "incratseq x (Suc n) = + (\q. max (of_rat(incratseq x n)) (x - 1/(n + 2)) < of_rat q \ + of_rat q < x)" + +lemma incratseq_0_gt [simp]: "x - 1 < of_rat(incratseq x 0)" +proof - + have "x - 1 < x" by simp + then have "\q. x - 1 < real_of_rat q \ real_of_rat q < x" + using of_rat_dense + by blast + then obtain q where " x - 1 < real_of_rat q \ real_of_rat q < x" + by force + then show ?thesis + by (auto intro: someI2) +qed + +lemma incratseq_0_less [simp]: "of_rat(incratseq x 0) < x" +proof - + have "x - 1 < x" by simp + then have "\q. x - 1 < real_of_rat q \ real_of_rat q < x" + using of_rat_dense + by blast + then obtain q where " x - 1 < real_of_rat q \ real_of_rat q < x" + by force + then show ?thesis + by (auto intro: someI2) +qed + +declare incratseq.simps [simp del] + +lemma incratseq_ub [simp]: + "real_of_rat (incratseq x n) < x" +proof (induction n) + case 0 + then show ?case by simp +next + case (Suc n) + then show ?case + proof (cases "real_of_rat (incratseq x n) \ x - 1/(n + 2)") + case True + then have "x - 1/(Suc(Suc n)) < x" by simp + then have "\q. x - 1 / real (Suc (Suc n)) + < real_of_rat q \ + real_of_rat q < x" + using of_rat_dense by blast + + then have "\a. (if True then x - 1 / real (n + 2) + else real_of_rat + (incratseq x n)) + < real_of_rat a \ + real_of_rat a < x" + by auto + then have "real_of_rat + (SOME q. + (if real_of_rat + (incratseq x n) + \ x - 1 / real (n + 2) + then x - 1 / real (n + 2) + else real_of_rat + (incratseq x n)) + < real_of_rat q \ + real_of_rat q < x) + < x" + using True by (fastforce intro: someI2_ex) + then show ?thesis + by (auto simp only: incratseq.simps max_def) + next + case False + then have "\a. real_of_rat (incratseq x n) + < real_of_rat a \ + real_of_rat a < x" + using Suc of_rat_dense by auto + then have "real_of_rat + (SOME q. + (if real_of_rat + (incratseq x n) + \ x - 1 / real (n + 2) + then x - 1 / real (n + 2) + else real_of_rat + (incratseq x n)) + < real_of_rat q \ + real_of_rat q < x) + < x" + using False by (fastforce intro: someI2_ex) + then show ?thesis + by (simp only: incratseq.simps max_def) + qed +qed + + +lemma incratseq_incseq [simp]: + "incratseq x n < incratseq x (Suc n)" +proof - + have "max (real_of_rat (incratseq x n)) (x - 1 / real (n + 2)) < x" + by simp + then have "\q. max (real_of_rat (incratseq x n)) + (x - 1 / real (n + 2)) + < real_of_rat q \ + real_of_rat q < x" + using of_rat_dense by blast + then have "incratseq x n + < (SOME q. + max (real_of_rat (incratseq x n)) + (x - 1 / real (n + 2)) + < real_of_rat q \ + real_of_rat q < x)" + by (fastforce intro: someI2_ex simp add: of_rat_less) + then show ?thesis + by (simp only: incratseq.simps) + qed + +lemma incratseq_lb [simp]: "x - 1/(Suc n) < real_of_rat (incratseq x n)" +proof (induction n) + case 0 + then show ?case by simp +next + case (Suc n) + then show ?case + proof (cases "real_of_rat (incratseq x n) \ x - 1/Suc(Suc n)") + case True + then have " 1 / Suc(Suc n) < 1 / Suc n" + using Suc.IH by auto + have "x - 1/Suc(Suc n) < x" + by simp + then have "\a. x - 1 /Suc(Suc n) < real_of_rat a \ real_of_rat a < x" + using of_rat_dense by blast + then have "x - 1 / real (Suc (Suc n)) + < real_of_rat + (SOME q. + (if real_of_rat (incratseq x n) \ x - 1 / Suc(Suc n) + then x - 1 / Suc(Suc n) else real_of_rat (incratseq x n)) + < real_of_rat q \ + real_of_rat q < x)" + using True by (auto intro: someI2_ex) + then show ?thesis using True + by (auto simp add: incratseq.simps(2) max_def) + next + case False + then show ?thesis + using incratseq_incseq + by (meson less_trans not_less of_rat_less) + qed +qed + +lemma incseq_incratseq [simp]: + "incseq (incratseq x)" +by (auto intro!: incseq_SucI less_imp_le) + + +lemma LIMSEQ_rat_real_ex: + "\r. incseq r \ (\n. of_rat (r n)) \ (x::real)" +proof - + have squeeze_left: + "\n. x - 1 / real (Suc n) + \ real_of_rat (incratseq x n) \ real_of_rat (incratseq x n) \ x" + using incratseq_lb less_imp_le + by (simp add: less_imp_le) + moreover have "(\n. x - 1 / real (Suc n)) \ x" + by (simp only: minus_real_def LIMSEQ_inverse_real_of_nat_add_minus + inverse_eq_divide [symmetric]) + ultimately have "(\n. real_of_rat (incratseq x n)) \ x" + using LIMSEQ_squeeze tendsto_const by fastforce + then show ?thesis + using incseq_incratseq by blast +qed + +lemma incseq_incseq_powrat: + "\ 1 \ a; incseq r \ \ incseq (\n. a pow\<^sub>\ (r n))" +by (metis (lifting) incseq_def powrat_le_mono) + +lemma ex_less_of_rat: "\r. (x :: 'a :: archimedean_field) < of_rat r" + using ex_less_of_int of_rat_of_int_eq by metis + +lemma powrat_incseq_bounded: + "\ 1 \ a; \n. r n < of_rat q; incseq r \ \ a pow\<^sub>\ (r n) \ a pow\<^sub>\ q" + by (simp add: less_imp_le powrat_le_mono) + +lemma Bseq_powrat_incseq: + assumes "1 \ a" + and "incseq r" + and "\n. of_rat(r n) \ (x :: 'a :: archimedean_field)" + shows "Bseq (\n. a pow\<^sub>\ (r n))" +proof - + from ex_less_of_rat + obtain q where xq: "x < of_rat q" by blast + then have "\n. 0 \ a pow\<^sub>\ r n \ a pow\<^sub>\ r n \ a pow\<^sub>\ of_rat q" + proof - + {fix m + have " r m < q" + using assms(3) le_less_trans of_rat_less xq by blast + then have "a pow\<^sub>\ r m \ a pow\<^sub>\ of_rat q" + by (simp add: assms(1) powrat_le_mono)} + then show ?thesis using less_imp_le [OF powrat_gt_zero] + using assms(1) by auto + qed + then show ?thesis + by (metis BseqI2' abs_of_nonneg real_norm_def) +qed + +lemma convergent_powrat_incseq: + "\ 1 \ a; incseq r; \n. r n \ x \ \ convergent (\n. a pow\<^sub>\ (r n))" + by (auto intro!: Bseq_monoseq_convergent + intro: incseq_imp_monoseq incseq_incseq_powrat Bseq_powrat_incseq [of _ _ x]) + +definition + incseq_of :: "real \ (nat \ rat)" where + "incseq_of x = (SOME r. incseq r \ (\n. of_rat (r n)) \ x)" + +lemma LIMSEQ_incseq_of [simp]: + "(\n. of_rat (incseq_of x n)) \ x" +by (auto intro: someI2_ex [OF LIMSEQ_rat_real_ex] simp add: incseq_of_def) + +lemma incseq_incseq_of [simp]: + "incseq (incseq_of x)" +by (auto intro: someI2_ex [OF LIMSEQ_rat_real_ex] simp add: incseq_of_def) + +lemma incseq_of_Suc [simp]: + "incseq_of x n \ incseq_of x (Suc n)" +by (metis incseq_def incseq_incseq_of le_SucI le_refl) + +lemma incseq_of_rat_incseq_of: + "incseq (\n. of_rat(incseq_of x n) :: 'a::linordered_field)" + by (meson incseq_def incseq_incseq_of of_rat_less_eq) + +lemma incseq_of_rat: + "incseq s \ incseq (\n. of_rat(s n) :: 'a :: linordered_field)" +by (auto simp add: incseq_def of_rat_less_eq) + +lemma incseq_rat_le_real: + "\ incseq s; (\n. of_rat (s n)) \ x \ \ of_rat (s n) \ (x::real)" +by (auto intro: incseq_le incseq_of_rat) + +lemma incseq_of_le_self: "of_rat(incseq_of x n) \ x" +by (auto intro!: incseq_rat_le_real LIMSEQ_incseq_of) + +lemma powrat_incseq_of_bounded: + "\ 1 \ a; x < of_rat r \ \ a pow\<^sub>\ (incseq_of x n) \ a pow\<^sub>\ r" + by (meson incseq_of_le_self le_less le_less_trans of_rat_less powrat_le_mono) + +lemma incseq_powrat_insec_of: + "1 \ a \ incseq (\n. a pow\<^sub>\ (incseq_of x n))" +by (auto intro: incseq_incseq_powrat) + +lemma Bseq_powrat_incseq_of: "1 \ a \ Bseq (\n. a pow\<^sub>\ (incseq_of x n))" +by (auto intro!: Bseq_powrat_incseq incseq_of_le_self) + +lemma convergent_powrat_incseq_of: "1 \ a \ convergent (\n. a pow\<^sub>\ (incseq_of x n))" +by (blast intro: Bseq_monoseq_convergent Bseq_powrat_incseq_of incseq_imp_monoseq + incseq_powrat_insec_of) + +text\We're now ready to define real exponentation.\ + +definition + powa :: "[real,real] \ real" (infixr "powa" 80) where + "a powa x = (THE y. (\n. a pow\<^sub>\ (incseq_of x n)) \ y)" + +text \real exponents.\ +definition + powreal :: "[real,real] \ real" (infixr "pow\<^sub>\" 80) where + "a pow\<^sub>\ x = (if 0 < a \ a < 1 then (inverse a) powa (-x) + else if a \ 1 then a powa x else 0)" + +lemma powreal_eq_powa: + "a \ 1 \ a pow\<^sub>\ x = a powa x" +by (simp add: powreal_def) + +lemma LIMSEQ_powrat_incseq_of_ex1: + "1 \ a \ \!y. (\n. a pow\<^sub>\ (incseq_of x n))\ y" + by (metis LIMSEQ_unique convergentD convergent_powrat_incseq_of) + +lemma LIMSEQ_powa: + "1 \ a \ (\n. a pow\<^sub>\ (incseq_of x n)) \ a powa x" +unfolding powa_def by (erule theI' [OF LIMSEQ_powrat_incseq_of_ex1]) + +lemma lemma_incseq_incseq_diff_inverse: + "incseq s \ incseq (\n. (s n :: rat) - 1/of_nat(Suc n))" + by (auto intro: diff_mono simp add: divide_inverse incseq_def) + +lemma lemma_incseq_diff_inverse_ub: + assumes "incseq s" + and "(\n. of_rat (s n)) \ x" +shows "of_rat(s n - 1/of_nat(Suc n)) < (x::real)" +proof - + have "real_of_rat (s n - 1 / rat_of_nat (Suc n)) < real_of_rat (s n)" + by (simp add: of_rat_diff) + then show ?thesis + by (meson assms incseq_rat_le_real linorder_not_less order_trans) +qed + +lemma lemma_LIMSEQ_incseq_diff_inverse: + assumes "(\n. of_rat (s n)) \ x" + shows " (\n. of_rat(s n - 1/of_nat(Suc n))) \ (x::real)" +proof - + have "(\x. real_of_rat (s x) - inverse (real (Suc x))) \ x" + using assms tendsto_diff [OF _ LIMSEQ_inverse_real_of_nat] + by simp + then show ?thesis + by (simp add: divide_inverse of_rat_diff of_rat_inverse of_rat_add) +qed + +lemma lemma_LIMSEQ_powrat_diff_inverse: + assumes "1 \ a" + and "(\n. a pow\<^sub>\ (s n))\ y" + shows "(\n. a pow\<^sub>\ (s n - 1/of_nat(Suc n))) \ y" +proof - + have "(\x. a pow\<^sub>\ (1 / rat_of_nat (Suc x))) \ 1" + using assms(1) LIMSEQ_powrat_inverse_of_nat + by (auto intro!: LIMSEQ_Suc simp only: divide_inverse mult_1_left) + then have " (\n. a pow\<^sub>\ s n / a pow\<^sub>\ (1 / rat_of_nat (Suc n))) \ y" + using assms(2) tendsto_divide by fastforce + then show ?thesis using powrat_diff assms(1) by auto +qed + +lemma lemma_LIMSEQ_powrat_diff_inverse2: + assumes "1 \ a" + and "(\n. a pow\<^sub>\ (s n - 1/of_nat(Suc n)))\ y" + shows "(\n. a pow\<^sub>\ (s n)) \ y" +proof - + have "(\x. a pow\<^sub>\ (1 / rat_of_nat (Suc x))) \ 1" + using assms(1) LIMSEQ_powrat_inverse_of_nat + by (auto intro!: LIMSEQ_Suc simp only: divide_inverse mult_1_left) + then have "(\x. a pow\<^sub>\ inverse (rat_of_nat (Suc x)) * + a pow\<^sub>\ (s x - inverse(rat_of_nat (Suc x)))) + \ 1 * y" + using assms(2) by (auto intro!: tendsto_mult simp only: inverse_eq_divide) + then show ?thesis using assms(1) + by (auto simp add: powrat_add [symmetric] ) +qed + +lemma lemma_seq_point_gt_ex: + "\ (\n. of_rat (r n)) \ (x::real); y < x \ + \ \(m::nat). y < of_rat(r m)" + by (metis convergent_def limI lim_le not_less) + +lemma lemma_seq_point_gt_ex2: + "\ (\n. of_rat (r n)) \ (x::real); of_rat y < x \ + \ (\m. y < r m)" + by (force dest: lemma_seq_point_gt_ex simp add: of_rat_less) + +primrec interlaced_index :: "(nat \ rat) \ (nat \ rat) \ nat \ nat" where + "interlaced_index r s 0 = 0" +| "interlaced_index r s (Suc n) = + (LEAST m. if odd n then r (interlaced_index r s n) < s m + else s (interlaced_index r s n) < r m)" + + +definition interlaced_seq :: "(nat \ rat) \ (nat \ rat) \ nat \ rat" where + "interlaced_seq r s n = (if odd n then r (interlaced_index r s n) + else s (interlaced_index r s n))" + +lemma incseq_interlaced_seq: + assumes "(\n. of_rat (r n)) \ (x::real)" + and "(\n. of_rat (s n)) \ (x::real)" + and "\n. of_rat (r n) < x" + and "\n. of_rat (s n) < x" + shows "incseq (interlaced_seq r s)" +proof - + {fix n + have "interlaced_seq r s n \ interlaced_seq r s (Suc n)" + using assms + by (auto intro: LeastI2_ex [OF lemma_seq_point_gt_ex2] + simp add: interlaced_seq_def)} + then show ?thesis + by (simp add: incseq_SucI) +qed + +lemma incseq_of_rat_interlaced_seq: + "\ (\n. of_rat (r n)) \ (x::real); + (\n. of_rat (s n)) \ (x::real); + \n. of_rat (r n) < x; \n. of_rat (s n) < x \ + \ incseq (\n. real_of_rat (interlaced_seq r s n))" + using incseq_interlaced_seq incseq_of_rat by blast + +lemma interlaced_seq_bounded: + "\ \n. of_rat (r n) < x; \n. of_rat (s n) < x\ + \ of_rat (interlaced_seq r s n) < x" +unfolding interlaced_seq_def by auto + + +lemma convergent_interlaced_seq: + assumes "(\n. of_rat (r n)) \ (x::real)" + and "(\n. of_rat (s n)) \ (x::real)" + and "\n. of_rat (r n) < x" + and "\n. of_rat (s n) < x" + shows "convergent (\n. real_of_rat (interlaced_seq r s n))" +proof - + have mono: "monoseq (\n. real_of_rat (interlaced_seq r s n))" + using assms incseq_interlaced_seq incseq_of_rat monoseq_iff by blast + {fix n + have "interlaced_seq r s 0 \ interlaced_seq r s n" + proof (induction n) + case 0 + then show ?case by simp + next + case (Suc n) + then show ?case + using assms incseq_def incseq_interlaced_seq by blast + qed} + moreover + {fix n + have "real_of_rat (interlaced_seq r s n) \ x" + by (simp add: assms(3,4) interlaced_seq_bounded le_less)} + ultimately have "Bseq (\n. real_of_rat (interlaced_seq r s n))" + by (metis decseq_bounded incseq_bounded mono monoseq_iff of_rat_less_eq) + then show ?thesis using mono + by (simp add: Bseq_monoseq_convergent) +qed + +lemma convergent_powrat_interlaced_seq: + "\ 1 \ a; (\n. of_rat (r n)) \ (x::real); + (\n. of_rat (s n)) \ (x::real); + \n. of_rat (r n) < x; \n. of_rat (s n) < x \ + \ convergent (\n. a pow\<^sub>\ (interlaced_seq r s n))" +by (auto intro: Bseq_monoseq_convergent incseq_interlaced_seq + interlaced_seq_bounded less_imp_le incseq_imp_monoseq + incseq_incseq_powrat Bseq_powrat_incseq [of _ _ x]) + +lemma LIMSEQ_even_odd_subseq_LIMSEQ: + assumes "(\n. (X (2 *n))) \ a" "(\n. (X (Suc(2 *n)))) \ a" + shows "X \ (a :: 'a::real_normed_vector)" +proof (auto simp add: LIMSEQ_iff) + fix r::real + assume e0: "r > 0" + obtain p where evenx: "\n\p. norm (X (2 * n) - a) < r" + using e0 assms(1) by (metis LIMSEQ_iff) + obtain q where oddx: "\n\q. norm (X (Suc (2 * n)) - a) < r" + using e0 assms(2) by (metis LIMSEQ_iff) + {fix n + assume max: "max (2 * p) (2 * q) \ n" + have "norm (X n - a) < r" + proof (cases "even n") + case True + then show ?thesis + using dvd_def evenx max + by auto + next + case False + then show ?thesis using oddx max + by (auto elim: oddE) + qed} + then show "\no. \n\no. norm (X n - a) < r" + by blast +qed + +(* Not proved before? *) +lemma incseqD2: "incseq r \ r m < r n \ m < n" + by (metis le_less mono_def not_le order.asym) + +lemma subseq_interlaced_index_even: + assumes "incseq r" + and "incseq s" + and "(\n. of_rat (r n)) \ (x::real)" + and "(\n. of_rat (s n)) \ (x::real)" + and "\n. of_rat (r n) < x" + and "\n. of_rat (s n) < x" + shows "strict_mono (\n. interlaced_index r s (2 * n))" +proof - + {fix n + have "interlaced_index r s (2 * n) + < (LEAST m. r (LEAST m. s (interlaced_index r s (2 * n)) < r m) < s m)" + proof (rule LeastI2_ex [of "\ m. s (interlaced_index r s (2 * n)) < r m"]) + show "\a. s (interlaced_index r s (2 * n)) < r a" + using assms(3) assms(6) lemma_seq_point_gt_ex2 by blast + next + fix m + assume inter_s: "s (interlaced_index r s (2 * n)) < r m" + show "interlaced_index r s (2 * n) < (LEAST ma. r m < s ma) " + proof (rule LeastI2_ex) + show "\a. r m < s a" + using assms(4) assms(5) lemma_seq_point_gt_ex2 by blast + next + fix k + assume "r m < s k" + then show "interlaced_index r s (2 * n) < k" + using inter_s assms(2) incseqD2 less_trans by blast + qed + qed} + then show ?thesis + by (simp add: strict_mono_Suc_iff) +qed + +lemma subseq_interlaced_index_odd: + assumes "incseq r" + and "incseq s" + and "(\n. of_rat (r n)) \ (x::real)" + and "(\n. of_rat (s n)) \ (x::real)" + and "\n. of_rat (r n) < x" + and "\n. of_rat (s n) < x" +shows "strict_mono (\n. interlaced_index r s (Suc (2 * n)))" +proof - + {fix n + have "(LEAST m. s (interlaced_index r s (2 * n)) < r m) + < (LEAST m. + s (LEAST m. r (LEAST m. s (interlaced_index r s (2 * n)) < r m) < s m) + < r m)" + proof (rule LeastI2_ex [of "(\ m. s (interlaced_index r s (2 * n)) < r m)"]) + show "\a. s (interlaced_index r s (2 * n)) < r a" + using assms(3) assms(6) lemma_seq_point_gt_ex2 by blast + next + fix m + assume inter_s: "s (interlaced_index r s (2 * n)) < r m" + show "m < (LEAST ma. s (LEAST ma. r m < s ma) < r ma)" + proof (rule LeastI2_ex [of "\ma. r m < s ma"]) + show "\a. r m < s a" + using assms(4) assms(5) lemma_seq_point_gt_ex2 by blast + next + fix k + assume rs: "r m < s k" + show "m < (LEAST ma. s k < r ma)" + proof (rule LeastI2_ex) + show "\a. s k < r a" + using assms(3) assms(6) lemma_seq_point_gt_ex2 by blast + next + fix l + assume "s k < r l" + then show "m < l" using rs + using assms(1) incseqD2 less_trans by blast + qed + qed + qed} + then show ?thesis + by (simp add: strict_mono_Suc_iff) +qed + +lemma interlaced_seq_even: + "interlaced_seq r s (2*n) = s (interlaced_index r s (2*n))" + by (simp add: interlaced_seq_def) + + +lemma interlaced_seq_odd: + "interlaced_seq r s (Suc (2*n)) = r (interlaced_index r s (Suc (2*n)))" + by (simp add: interlaced_seq_def) + +lemma powa_indep_incseq_of: + assumes "1 \ a" + and "incseq r" + and "incseq s" + and "(\n. real_of_rat (r n)) \ x" + and "(\n. real_of_rat (s n)) \ x" + and "(\n. a pow\<^sub>\ (r n)) \ y" + and "(\n. a pow\<^sub>\ (s n)) \ z" +shows "y = z" +proof - + have rx: "(\n. real_of_rat (r n - 1 / rat_of_nat (Suc n))) \ x" + using assms(4) lemma_LIMSEQ_incseq_diff_inverse by blast + have sx: "(\n. real_of_rat (s n - 1 / rat_of_nat (Suc n))) \ x" + using assms(5) lemma_LIMSEQ_incseq_diff_inverse by blast + have "convergent + (\n. a pow\<^sub>\ + interlaced_seq (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) n)" + using convergent_powrat_interlaced_seq + [of _ "(\n. r n - 1 / rat_of_nat (Suc n))" + _ "(\n. s n - 1 / rat_of_nat (Suc n))"] + assms(1,2,3,4,5) lemma_incseq_diff_inverse_ub + lemma_LIMSEQ_incseq_diff_inverse by blast + then obtain L + where converge: "(\n. a pow\<^sub>\ + interlaced_seq (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) n) \ L" + using convergent_def by force + then have "((\n. a pow\<^sub>\ + interlaced_seq (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) n) \ (\n. 2 * n)) \ L" + by (simp add: LIMSEQ_subseq_LIMSEQ strict_mono_Suc_iff) + moreover have "((\n. a pow\<^sub>\ + interlaced_seq (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) n) \ (\n. Suc (2 * n))) \ L" + using converge by (simp add: LIMSEQ_subseq_LIMSEQ strict_mono_Suc_iff) + moreover have "((\n. a pow\<^sub>\ (r n - 1 / rat_of_nat (Suc n))) \ + (\n. interlaced_index (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) (Suc (2 * n)))) \ y" + proof - + from rx sx + have "strict_mono + (\n. interlaced_index (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) (Suc (2 * n)))" + using subseq_interlaced_index_odd lemma_incseq_incseq_diff_inverse + lemma_incseq_diff_inverse_ub assms by blast + moreover have "(\n. a pow\<^sub>\ (r n - 1 / rat_of_nat (Suc n))) \ y" + using assms(1) assms(6) lemma_LIMSEQ_powrat_diff_inverse by blast + ultimately show ?thesis + using LIMSEQ_subseq_LIMSEQ by blast + qed + moreover have "((\n. a pow\<^sub>\ (s n - 1 / rat_of_nat (Suc n))) \ + (\n. interlaced_index (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) (2 * n))) + \ z" + proof - + from rx sx + have "strict_mono + (\n. interlaced_index (\n. r n - 1 / rat_of_nat (Suc n)) + (\n. s n - 1 / rat_of_nat (Suc n)) (2 * n))" + using subseq_interlaced_index_even lemma_incseq_incseq_diff_inverse + lemma_incseq_diff_inverse_ub assms by blast + moreover have "(\n. a pow\<^sub>\ (s n - 1 / rat_of_nat (Suc n))) \ z" + using assms(1) assms(7) lemma_LIMSEQ_powrat_diff_inverse by blast + ultimately show ?thesis + using LIMSEQ_subseq_LIMSEQ by blast + qed + ultimately show ?thesis + by (force dest: LIMSEQ_unique simp only: o_def interlaced_seq_even interlaced_seq_odd) +qed + +lemma powa_indep_incseq_of': + "\ 1 \ a; incseq r; + (\n. real_of_rat (r n)) \ x; + (\n. a pow\<^sub>\ (r n)) \ y \ + \ (\n. a pow\<^sub>\ (incseq_of x n)) \ y" + using powa_indep_incseq_of [of a r "incseq_of x"] LIMSEQ_powa + by fastforce + +(* Similar theorems as above, but specialized to incseq_of *) +lemma lemma_incseq_incseq_of_diff_inverse: + "incseq (\n. incseq_of x n - 1/of_nat(Suc n))" +by (blast intro: lemma_incseq_incseq_diff_inverse incseq_incseq_of) + +lemma lemma_incseq_of_diff_inverse_ub: + "of_rat(incseq_of x n - 1/of_nat(Suc n)) < x" +by (blast intro: lemma_incseq_diff_inverse_ub incseq_incseq_of LIMSEQ_incseq_of) + + lemma lemma_LIMSEQ_incseq_of_diff_inverse: + "(\n. of_rat(incseq_of x n - 1/of_nat(Suc n))) \ x" +by (blast intro: lemma_LIMSEQ_incseq_diff_inverse incseq_incseq_of LIMSEQ_incseq_of) + +lemma powa_add: + assumes "1 \ a" + shows "a powa (x + y) = a powa x * a powa y" +proof - + obtain k where 1: "(\n. a pow\<^sub>\ incseq_of y n) \ k" + using LIMSEQ_powrat_incseq_of_ex1 assms by blast + moreover obtain l where 2: "(\n. a pow\<^sub>\ incseq_of x n) \ l" + using LIMSEQ_powrat_incseq_of_ex1 assms by blast + ultimately have "(\n. a pow\<^sub>\ (incseq_of x n + incseq_of y n)) \ l * k" + using assms by (auto intro: tendsto_mult simp add: powrat_add ) + moreover + have "incseq (\n. incseq_of x n + incseq_of y n)" + by (force intro: incseq_SucI add_mono) + moreover have "(\n. real_of_rat (incseq_of x n + incseq_of y n)) \ x + y" + by (auto simp add: of_rat_add intro: tendsto_add LIMSEQ_incseq_of) + ultimately have "(\n. a pow\<^sub>\ incseq_of (x + y) n) \ l * k" + using assms powa_indep_incseq_of' by blast + then show ?thesis using powa_def 1 2 + by (metis Lim_def limI) +qed + +lemma real_inverse_ge_one_lemma: + "\ 0 < (a::real); a < 1 \ \ inverse a \ 1" +by (metis less_eq_real_def one_le_inverse_iff) + +lemma real_inverse_gt_one_lemma: + "\ 0 < (a::real); a < 1 \ \ inverse a > 1" +by (metis one_less_inverse_iff) + +lemma real_inverse_bet_one_one_lemma: + "1 < (a::real) \ 0 < inverse a \ inverse a < 1" + by (metis inverse_less_1_iff inverse_positive_iff_positive + le_less_trans zero_le_one) + +lemma powreal_add: + "a pow\<^sub>\ (x + y) = a pow\<^sub>\ x * a pow\<^sub>\ y" + by (metis minus_add_distrib mult_zero_right powa_add + powreal_def real_inverse_ge_one_lemma) + +lemma powa_one_eq_one [simp]: "1 powa a = 1" +proof - + have "(\n. 1 pow\<^sub>\ incseq_of a n) \ 1" + by simp + then show ?thesis + by (metis Lim_def limI powa_def) +qed + +lemma powreal_one_eq_one [simp]: "1 pow\<^sub>\ a = 1" +by (simp add: powreal_def) + +lemma powa_zero_eq_one [simp]: "a \ 1 \ a powa 0 = 1" +by (auto intro!: the1_equality LIMSEQ_powrat_incseq_of_ex1 + intro: powa_indep_incseq_of' [of a "\n. 0"] incseq_SucI + simp add: powa_def) + +lemma powreal_zero_eq_one [simp]: "a > 0 \ a pow\<^sub>\ 0 = 1" +by (auto dest: real_inverse_ge_one_lemma simp add: powreal_def) + +lemma powr_zero_eq_one_iff [simp]: "x pow\<^sub>\ 0 = (if x \ 0 then 0 else 1)" + using powreal_def powreal_zero_eq_one by force + +lemma powa_one_gt_zero [simp]: "1 \ a \ a powa 1 = a" +by (auto intro!: LIMSEQ_powrat_incseq_of_ex1 the1_equality + powa_indep_incseq_of' [of a "\n. 1"] incseq_SucI simp add: powa_def) + +lemma powa_minus_one: + assumes "1 \ a" shows "a powa -1 = inverse a" +proof - + have "(\n. a pow\<^sub>\ - 1) \ inverse a" using assms + by (simp add: powrat_minus) + then have "(\n. a pow\<^sub>\ incseq_of (- 1) n) \ inverse a" + using powa_indep_incseq_of' [of a "\n. -1"] assms + by simp + then show ?thesis using powa_def + by (metis Lim_def limI) +qed + +lemma powreal_minus_one: "0 \ a \ a pow\<^sub>\ -1 = inverse a" + by (simp add: powa_minus_one powreal_def real_inverse_ge_one_lemma) + +lemma powreal_one [simp]: "a \ 0 \ a pow\<^sub>\ 1 = a" + by (simp add: powa_minus_one powreal_def real_inverse_ge_one_lemma) + +lemma powa_gt_zero: + assumes "a \ 1" + shows "a powa x > 0" +proof - + obtain y where 1: "(\n. a pow\<^sub>\ incseq_of x n) \ y" + using LIMSEQ_powrat_incseq_of_ex1 assms by blast + moreover have "incseq (\n. a pow\<^sub>\ incseq_of x n)" + by (simp add: assms incseq_powrat_insec_of) + ultimately have "y > 0" + using assms incseq_le powrat_gt_zero + by (metis less_le_trans zero_less_one) + then show ?thesis using powa_def 1 + by (metis Lim_def limI) +qed + +lemma powreal_gt_zero: "a > 0 \ a pow\<^sub>\ x > 0" + by (auto dest: powa_gt_zero real_inverse_ge_one_lemma + simp add: powreal_def not_less) + +lemma powreal_not_zero: "a > 0 \ a pow\<^sub>\ x \ 0" +by (metis order_less_imp_not_eq powreal_gt_zero) + +lemma powreal_minus: + "a pow\<^sub>\ -x = inverse (a pow\<^sub>\ x)" +proof (cases "a < 0") + case True + then show ?thesis + using powreal_def by force +next + case False + then have "a pow\<^sub>\ (x + -x) = a pow\<^sub>\ x * a pow\<^sub>\ -x" + using powreal_add by presburger + then show ?thesis + using inverse_unique powreal_def powreal_zero_eq_one + by fastforce +qed + +lemma powreal_minus_base_ge_one: + "a pow\<^sub>\ (-x) = (inverse a) pow\<^sub>\ x" +using one_le_inverse_iff powreal_def by auto + +lemma powreal_inverse: + "inverse (a pow\<^sub>\ x) = (inverse a) pow\<^sub>\ x" + using powreal_minus powreal_minus_base_ge_one + by presburger + +lemma powa_minus:"a \ 1 \ a powa (-x) = inverse (a powa x)" + by (metis powreal_eq_powa powreal_minus) + +lemma powa_mult_base: + assumes "1 \ a" and "1 \ b" + shows "(a * b) powa x = (a powa x) * (b powa x)" +proof - + obtain x' where lim_a: "(\n. a pow\<^sub>\ incseq_of x n) \ x'" + using LIMSEQ_powrat_incseq_of_ex1 assms(1) by blast + then have lim_a_eq: "(THE y. (\n. a pow\<^sub>\ incseq_of x n) \ y) = x'" + by (metis Lim_def limI) + obtain y' where lim_b: "(\n. b pow\<^sub>\ incseq_of x n) \ y'" + using LIMSEQ_powrat_incseq_of_ex1 assms(2) by blast + then have lim_b_eq: "(THE y. (\n. b pow\<^sub>\ incseq_of x n) \ y) = y'" + by (metis Lim_def limI) + have lim_ab: "(\n. (a * b) pow\<^sub>\ incseq_of x n) \ x' * y'" + using lim_a lim_b by (auto dest: tendsto_mult simp add: powrat_mult_base) + then have "(THE y. (\n. (a * b) pow\<^sub>\ incseq_of x n) \ y) = x' * y'" + by (metis Lim_def limI) + then show ?thesis + by (simp add: lim_a_eq lim_b_eq powa_def) + qed + + +lemma powreal_mult_base_lemma1: + "\ 1 \ a; 1 \ b \ + \ (a * b) pow\<^sub>\ x = (a pow\<^sub>\ x) * (b pow\<^sub>\ x)" +by (metis mult.left_neutral mult_mono' powa_mult_base powreal_eq_powa zero_le_one) + +lemma powreal_mult_base_lemma2: + assumes "1 \ a" + and "0 < b" + and "b < 1" +shows "(a * b) pow\<^sub>\ x = (a pow\<^sub>\ x) * (b pow\<^sub>\ x)" +proof (cases "a * b < 1") + case True + assume ab1: "a * b < 1" + have "a * b > 0" + using assms(1) assms(2) by auto + then have inv_ab1: "1 \ inverse (a * b)" + using ab1 real_inverse_ge_one_lemma by blast + then have "(a * (inverse a * inverse b)) powa - x = + a powa - x * (inverse a * inverse b) powa - x" + by (simp add: assms(1) powa_mult_base) + then have "(inverse b) powa - x = + a powa - x * (inverse a * inverse b) powa - x" + using assms(1) by (simp add: mult.assoc [symmetric]) + then have "(inverse a * inverse b) powa - x = a powa x * inverse b powa - x" + by (simp add: mult.assoc [symmetric] powa_add [symmetric] assms(1)) + then show ?thesis + using ab1 assms powreal_def by auto +next + case False + have inv_b: "inverse b \ 1" + by (simp add: assms real_inverse_ge_one_lemma) + assume "\ a * b < 1" + then have "a * b \ 1" by simp + then have "(a * b * inverse b) powa x = (a * b) powa x * inverse b powa x" + by (simp add: assms(2) assms(3) powa_mult_base real_inverse_ge_one_lemma) + then have "a powa x = (a * b) powa x * inverse b powa x" + using assms(2) by (auto simp add: mult.assoc) + then have "(a * b) powa x = a powa x * inverse b powa - x" + by (simp add: mult.assoc powa_add [symmetric] assms inv_b) + then show ?thesis + using False assms powreal_def by auto +qed + +lemma powreal_mult_base_lemma3: + assumes "0 < a" + and "a < 1" + and "0 < b" + and "b < 1" +shows "(a * b) pow\<^sub>\ x = (a pow\<^sub>\ x) * (b pow\<^sub>\ x)" +proof - + have "a * b < 1" using assms + by (metis less_trans mult.left_neutral mult_less_cancel_right_disj) + moreover have "(inverse a * inverse b) powa - x = + inverse a powa - x * inverse b powa - x" + by (simp add: assms powa_mult_base real_inverse_ge_one_lemma) + ultimately show ?thesis using powreal_def assms by simp +qed + +lemma powreal_mult_base: + assumes "0 \ a" and "0 \ b" + shows "(a * b) pow\<^sub>\ x = (a pow\<^sub>\ x) * (b pow\<^sub>\ x)" +proof (cases "a \ 1 \ b \ 1") + case True + then show ?thesis + by (simp add: powreal_mult_base_lemma1) +next + case False + then show ?thesis + using powreal_mult_base_lemma3 powreal_mult_base_lemma2 assms + by (smt (verit) mult.commute mult_minus_left powreal_def) +qed + +lemma incseq_le_all: "incseq X \ X \ L \ \n. X n \ (L::real)" +by (metis incseq_le) + +lemma powa_powrat_eq: + assumes "a \ 1" shows "a powa (of_rat q) = a pow\<^sub>\ q" +proof - + have "(\n. a pow\<^sub>\ incseq_of (real_of_rat q) n) \ a pow\<^sub>\ q" + using powa_indep_incseq_of' assms by fastforce + then show ?thesis using powa_def + by (metis Lim_def limI) +qed + + +lemma realpow_powrat_eq: "a > 0 \ a pow\<^sub>\ (of_rat q) = a pow\<^sub>\ q" +proof - + assume a1: "0 < a" + then have "\ a < 1 \ 1 \ inverse a" + using real_inverse_ge_one_lemma by blast + then show ?thesis + using a1 + by (metis inverse_inverse_eq not_le powa_powrat_eq + powrat_inverse powreal_eq_powa powreal_inverse) +qed + +(* Move to NthRoot.thy *) +lemma LIMSEQ_real_root: + "\ X \ a; m > 0 \ \ (\n. root m (X n)) \ (root m a)" +by (metis isCont_tendsto_compose [where g="\x. root m x"] isCont_real_root) + +lemma powa_powrat_lemma1: + assumes "a \ 1" and "p \ 0" + shows "(a powa x) pow\<^sub>\ p = (a powa (x * of_rat p))" +proof - + obtain y where lim_a: "(\n. a pow\<^sub>\ incseq_of x n) \ y" + using LIMSEQ_powrat_incseq_of_ex1 assms(1) by blast + then have the_lim: "(THE y. (\n. a pow\<^sub>\ incseq_of x n) \ y) = y" + by (metis Lim_def limI) + then have "(\n. (a pow\<^sub>\ incseq_of x n) pow\<^sub>\ p) \ y pow\<^sub>\ p" + using LIMSEQ_powrat_base assms(1) lim_a powa_def powa_gt_zero by auto + then have lim_ap: "(\n. a pow\<^sub>\ (incseq_of x n * p)) \ y pow\<^sub>\ p" + using assms(1) powrat_mult by auto + then have "(\n. a pow\<^sub>\ incseq_of (x * real_of_rat p) n) \ y pow\<^sub>\ p" + proof - + have "incseq (\n. incseq_of x n * p)" + by (simp add: assms(2) incseq_SucI mult_right_mono) + moreover have "(\n. real_of_rat (incseq_of x n * p)) \ x * real_of_rat p" + by (auto intro!: tendsto_mult simp add: of_rat_mult) + ultimately show ?thesis + using assms(1) lim_ap powa_indep_incseq_of' by blast + qed + then show ?thesis using powa_def + by (metis Lim_def limI the_lim) + qed + +lemma powa_powrat_lemma2: + assumes "a \ 1" and "p < 0" + shows "(a powa x) pow\<^sub>\ p = (a powa (x * of_rat p))" +proof - + have "(a powa x) pow\<^sub>\ - p = a powa (x * real_of_rat (- p))" + by (simp add: assms(1) assms(2) less_imp_le powa_powrat_lemma1) + then have "inverse ((a powa x) pow\<^sub>\ p) = inverse (a powa (x * real_of_rat p))" + by (simp add: assms(1) of_rat_minus powa_minus powrat_minus) + then show ?thesis by simp +qed + +lemma powa_powrat_lemma: + "a \ 1 \ (a powa x) pow\<^sub>\ p = (a powa (x * of_rat p))" +by (metis linorder_not_less powa_powrat_lemma1 powa_powrat_lemma2) + +(* Didn't we use to have something similar proved? *) +lemma LIMSEQ_iff2: + fixes L :: "'a::metric_space" + shows "(X \ L) = (\m::nat>0. \no. \n\no. dist (X n) L < inverse m)" +proof + assume "X \ L" + then show "\m>0. \no. \n\no. dist (X n) L < inverse (real m)" + by (auto simp add: LIMSEQ_def) +next + assume assm: "\m>0. \no. \n\no. dist (X n) L < inverse (real m)" + {fix r + assume "(r::real) > 0" + then have " \no. \n\no. dist (X n) L < r" + using assm + by (metis ex_inverse_of_nat_less less_trans) + } + then show "X \ L" + by (simp add: metric_LIMSEQ_I) +qed + +lemma LIM_def2: + "f \a\ L = (\m::nat>0. \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < inverse m)" + for a :: "'a::metric_space" and L :: "'b::metric_space" +proof + assume "f \a\ L" + then show "\m::nat>0. \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < inverse m" + by (auto simp add: LIM_def) +next + assume assm: "\m::nat>0. \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < inverse m" + {fix r + assume r0: "(r::real) > 0" + then obtain n where "inverse (real (Suc n)) < r" + using reals_Archimedean by blast + then have "\s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" + using assm r0 by (metis order_less_trans zero_less_Suc) + } + then show "f \a\ L" + by (simp add: metric_LIM_I) +qed + +lemma powa_ge_one: + assumes "a \ 1" + and "x \ 0" + shows "a powa x \ 1" +proof - + obtain y where lima: "(\n. a pow\<^sub>\ incseq_of x n) \ y" + using LIMSEQ_powrat_incseq_of_ex1 assms(1) by blast + moreover have "incseq (\n. a pow\<^sub>\ incseq_of x n)" + using assms(1) incseq_powrat_insec_of by blast + moreover have "\n. a pow\<^sub>\ incseq_of x n \ y" + using incseq_le_all using calculation by blast + ultimately have "1 \ y" + using assms LIMSEQ_incseq_of LIMSEQ_powa LIMSEQ_powrat_incseq_of_ex1 + lemma_seq_point_gt_ex2 powa_zero_eq_one powrat_ge_one + by (metis less_le of_rat_0 less_eq_real_def xt1(6)) + then show ?thesis + using LIMSEQ_powa LIMSEQ_unique assms(1) lima by blast +qed + +lemma powreal_ge_one: "a \ 1 \ x \ 0 \ a pow\<^sub>\ x \ 1" +by (simp add: powreal_def powa_ge_one) + +lemma powreal_ge_one2: + "\ 0 < a; a < 1; x \ 0 \ \ a pow\<^sub>\ x \ 1" + by (simp add: powa_ge_one powreal_def real_inverse_ge_one_lemma) + +lemma inverse_of_real_nat_of_rat_of_nat: + "inverse (real_of_nat n) = of_rat(inverse (of_nat n))" +by (metis Ratreal_def of_rat_of_nat_eq real_inverse_code) + +lemma LIMSEQ_powa_inverse_of_nat: + "a \ 1 \ (\n. a powa inverse (real_of_nat n)) \ 1" + by (simp add: inverse_of_real_nat_of_rat_of_nat powa_powrat_eq + LIMSEQ_powrat_inverse_of_nat) + +lemma incseq_of_le_mono: + assumes "r \ s" + shows "\N. \n\N. incseq_of r n \ incseq_of s n" +proof - + have "r = s \ r < s" using assms by force + then show ?thesis + proof + assume "r = s" then show ?thesis by simp + next + assume rs: "r < s" + then obtain m where less_incseq: "r < real_of_rat (incseq_of s m)" + using LIMSEQ_incseq_of lemma_seq_point_gt_ex by blast + moreover have "\n. real_of_rat (incseq_of r n) \ r" + using incseq_of_le_self by simp + ultimately have "\n. real_of_rat (incseq_of r n) < real_of_rat (incseq_of s m)" + using le_less_trans by blast + then have incrs: "\n. incseq_of r n \ incseq_of s m" + by (simp add: less_imp_le of_rat_less) + then show ?thesis + by (meson incseq_def incseq_incseq_of order_trans) + qed +qed + +lemma powa_le_mono: + assumes "r \ s" + and "a \ 1" + shows "a powa r \ a powa s" +proof - + obtain y where "(\n. a pow\<^sub>\ incseq_of s n) \ y" + using LIMSEQ_powrat_incseq_of_ex1 assms(2) by blast + moreover obtain x where "(\n. a pow\<^sub>\ incseq_of r n) \ x" + using LIMSEQ_powrat_incseq_of_ex1 assms(2) by blast + moreover have "\N. \n\N. a pow\<^sub>\ incseq_of r n \ a pow\<^sub>\ incseq_of s n" + by (meson assms(1) assms(2) incseq_of_le_mono powrat_le_mono) + moreover have "x \ y" + using LIMSEQ_le calculation by blast + ultimately show ?thesis + by (metis Lim_def limI powa_def) +qed + +lemma powreal_le_mono: + "\ r \ s; a \ 1 \ \ a pow\<^sub>\ r \ a pow\<^sub>\ s" +by (metis powa_le_mono powreal_eq_powa) + +lemma powreal_le_anti_mono: + "\ r \ s; 0 < a; a < 1 \ \ a pow\<^sub>\ r \ a pow\<^sub>\ s" + by (simp add: powa_le_mono powreal_def real_inverse_ge_one_lemma) + +lemma powreal_less_cancel: + "\ a pow\<^sub>\ r < a pow\<^sub>\ s; a \ 1 \ \ r < s" +by (metis less_le_not_le linorder_not_less powreal_eq_powa powreal_le_mono) + +lemma powa_less_mono: + assumes "r < s" and "a > 1" + shows "a powa r < a powa s" +proof - + obtain q where "r < real_of_rat q" "real_of_rat q < s" + using assms(1) of_rat_dense by blast + moreover obtain qa where "real_of_rat q < real_of_rat qa" "real_of_rat qa < s" + using of_rat_dense calculation(2) by blast + ultimately show ?thesis using assms + by (metis (full_types) antisym less_eq_real_def less_imp_not_eq2 powa_le_mono + powa_powrat_eq powrat_inject_exp) + qed + +lemma powreal_less_anti_mono: + assumes "r < s" + and "0 < a" + and "a < 1" +shows "a pow\<^sub>\ r > a pow\<^sub>\ s" +proof - + have "inverse a > 1" + by (simp add: assms(2, 3) one_less_inverse_iff) + moreover have "inverse a powa r < inverse a powa s" + using assms(1) powa_less_mono + using calculation by blast + ultimately show ?thesis + using powreal_eq_powa powreal_inverse powreal_le_anti_mono powreal_less_cancel + by (metis assms(2,3) le_less less_irrefl) +qed + +lemma powreal_less_mono: + "\ r < s; a > 1 \ \ a pow\<^sub>\ r < a pow\<^sub>\ s" +by (simp add: powreal_def powa_less_mono) + +lemma powa_le_cancel: + "\ a powa r \ a powa s; a > 1 \ \ r \ s" +by (metis not_le powa_less_mono) + +lemma powreal_le_cancel: + "\ a pow\<^sub>\ r \ a pow\<^sub>\ s; a > 1 \ \ r \ s" +by (metis not_le powreal_less_mono) + +lemma powreal_less_cancel_iff [simp]: + "1 < a \ (a pow\<^sub>\ r < a pow\<^sub>\ s) = (r < s)" +by (metis less_imp_le powreal_less_cancel powreal_less_mono) + +lemma powreal_le_cancel_iff [simp]: + "1 < a \ (a pow\<^sub>\ r \ a pow\<^sub>\ s) = (r \ s)" +by (simp add: linorder_not_less [symmetric]) + +lemma powreal_inject_exp1 [simp]: + "1 < a \ (a pow\<^sub>\ r = a pow\<^sub>\ s) = (s = r)" +by (metis antisym_conv3 less_irrefl powreal_less_mono) + +lemma powreal_eq_one_iff [simp]: + "a pow\<^sub>\ x = 1 \ x = 0" if "a > 1" + using powreal_inject_exp1 that by fastforce + +lemma powreal_inject_base_less_one [simp]: + "0 < a \ a < 1 \ (a pow\<^sub>\ r = a pow\<^sub>\ s) = (s = r)" +by (metis linorder_neq_iff order_less_imp_not_eq2 powreal_less_anti_mono) + +lemma powreal_inject [simp]: + "0 < a \ a \ 1 \ (a pow\<^sub>\ r = a pow\<^sub>\ s) = (s = r)" + using powreal_inject_base_less_one by fastforce + +lemma powreal_gt_one: "a > 1 \ x > 0 \ a pow\<^sub>\ x > 1" +by (metis less_eq_real_def powa_less_mono powa_zero_eq_one powreal_eq_powa) + +lemma isCont_powa_exponent_at_zero: + assumes "a > 1" shows "isCont (\x. a powa x) 0" +proof - + {fix m + assume m0: "(m::nat) > 0" + have lim1: "(\n. a powa inverse (real n)) \ 1" + by (simp add: LIMSEQ_powa_inverse_of_nat assms less_imp_le) + then have "\no. \n\no. \a powa inverse (real n) - 1\ < inverse (real m)" + using lim1 LIMSEQ_iff2 dist_real_def m0 by metis + then obtain "no" + where pow1: "\n\no. \a powa inverse (real n) - 1\ < inverse (real m)" + by blast + + have lim2: "(\x. inverse (a powa inverse (real x))) \ 1" + using tendsto_inverse lim1 by fastforce + then have "\k. \n\k. \inverse (a powa inverse (real n)) - 1\ < inverse (real m)" + using LIMSEQ_iff2 dist_real_def m0 by metis + then obtain "k" + where pow2: "\n\k. \inverse (a powa inverse (real n)) - 1\ < inverse (real m)" + by blast + have "\s>0. \x. x \ 0 \ \x\ < s \ \(a powa x) - 1\ < inverse (real m)" + proof - + let ?d = "min (inverse (Suc no)) (inverse (Suc k))" + {fix x + assume "abs x < ?d" + then have x_bounds: "- inverse (of_nat(Suc k)) < x" "x < inverse (of_nat(Suc no))" + by linarith+ + then have "a powa - inverse (of_nat (Suc k)) < a powa x" + using assms powa_less_mono by blast + moreover have "- inverse (real m) < a powa - inverse (real (Suc k)) - 1" + using assms pow2 powa_minus + by (metis minus_diff_eq neg_less_iff_less abs_less_iff lessI less_imp_le) + ultimately have lo: "- inverse(real m) < a powa x - 1" + by linarith + have "a powa x < a powa inverse (of_nat(Suc no))" + using assms powa_less_mono x_bounds(2) by blast + moreover have "a powa inverse (real (Suc no)) - 1 < inverse (real m)" + using assms pow1 by (metis less_imp_le abs_less_iff lessI) + ultimately have "a powa x - 1 < inverse(real m)" + by linarith + then have "abs (a powa x - 1) < inverse(real m)" + using lo by linarith} + then show ?thesis + by (metis inverse_positive_iff_positive min_less_iff_conj of_nat_0_less_iff zero_less_Suc) + qed} + then have "\m>0. \s>0. \x. x \ 0 \ \x\ < s \ \a powa x - 1\ < inverse (real m)" + by blast + then show ?thesis + by (auto simp add: assms isCont_def LIM_def2 dist_real_def less_imp_le) +qed + +lemma LIM_powa_exponent_at_zero: "1 < a \ (\h. a powa h) \0\ 1" +by (metis isCont_def isCont_powa_exponent_at_zero less_eq_real_def powa_zero_eq_one) + +lemma isCont_powa_exponent_gt_one: + assumes "a > 1" + shows "isCont (\x. a powa x) x" +proof - + have "(\h. a powa x * a powa h) \0\ a powa x" + using LIM_powa_exponent_at_zero assms tendsto_mult_left by fastforce + then have "(\h. a powa (x + h)) \0\ a powa x" + using assms powa_add by auto + then have "(powa) a \x\ a powa x" + using LIM_offset_zero_cancel by blast + then show ?thesis + using isCont_def by blast +qed + +lemma isCont_powreal_exponent_gt_one: + "a > 1 \ isCont (\x. a pow\<^sub>\ x) x" +by (metis ext isCont_powa_exponent_gt_one less_eq_real_def powreal_eq_powa) + +lemma isCont_powreal_exponent_less_one: + assumes "0 < a" + and "a < 1" + shows "isCont (\x. a pow\<^sub>\ x) x" +proof - + have "1 < inverse a" + by (simp add: assms one_less_inverse) + then have "isCont ((pow\<^sub>\) (inverse a)) x" + by (simp add: isCont_powreal_exponent_gt_one) + then have "isCont (\x. inverse (inverse a pow\<^sub>\ x)) x" + using assms(1) continuous_at_within_inverse powreal_not_zero by fastforce + then show ?thesis + using assms(1) powreal_inverse by auto +qed + +lemma isCont_powreal_exponent: + assumes a_gt_0: "0 < a" shows "isCont (\x. a pow\<^sub>\ x) x" +proof cases + assume "a > 1" then show ?thesis + using isCont_powreal_exponent_gt_one by blast + next + assume "\ a > 1" + then have "a < 1 \ a = 1" by auto + then show ?thesis + proof + assume "a < 1" then show ?thesis + using a_gt_0 isCont_powreal_exponent_less_one by blast + next + assume "a =1" then show ?thesis by simp + qed +qed + +(* A little diversion *) +lemma real_of_rat_abs: + "real_of_rat(abs a) = abs(of_rat a)" +by (simp add: abs_if of_rat_minus) + +lemma isCont_powrat_exponent: + assumes "0 < a" shows "isCont (\x. a pow\<^sub>\ x) x" +proof - + have isCont_of_rat: "isCont real_of_rat x" using isCont_def LIM_def dist_real_def + by (metis dist_commute of_rat_diff rat_dist_def real_of_rat_abs) + have "isCont ((pow\<^sub>\) a) (real_of_rat x)" using assms + by (simp add: isCont_powreal_exponent) + then have "isCont (\x. a pow\<^sub>\ real_of_rat x) x" using isCont_of_rat + using isCont_o2 by blast + then show ?thesis using realpow_powrat_eq assms by simp +qed + + +lemma LIMSEQ_powrat_exponent: + "\ X \ x; a > 0 \ \ (\n. a pow\<^sub>\ (X n)) \ a pow\<^sub>\ x" +by (metis isCont_tendsto_compose isCont_powrat_exponent) + +(* Now, back to business *) +lemma powa_mult: + assumes "1 \ a" and "0 \ x" + shows "(a powa x) powa y = a powa (x * y)" +proof (cases) + assume "a \ 1" + then have a_gt_1: "a > 1" using assms(1) by simp + have "a powa x \ 1" using assms powa_ge_one by blast + then have lim1: "(\n. a powa (x * real_of_rat (incseq_of y n))) \ (a powa x) powa y" + using LIMSEQ_powa [of "a powa x" y] powa_powrat_lemma assms(1) by simp + have "isCont ((powa) a) (x * y)" using isCont_powa_exponent_gt_one a_gt_1 by blast + moreover have "(\n. x * real_of_rat (incseq_of y n)) \ x * y" + by (simp add: tendsto_mult_left) + ultimately have "(\n. a powa (x * real_of_rat (incseq_of y n))) \ a powa (x * y)" + using isCont_tendsto_compose by blast + then show ?thesis + using LIMSEQ_unique lim1 by blast +next + assume "\ a \ 1" + then show ?thesis + by auto +qed + +lemma powreal_mult1: + "\ 1 \ a; 0 \ x \ \ (a pow\<^sub>\ x) pow\<^sub>\ y = a pow\<^sub>\ (x * y)" +by (metis powa_mult powreal_eq_powa powreal_ge_one) + +lemma powreal_mult2: + assumes "0 < a" and "a < 1" and "0 \ x" + shows "(a pow\<^sub>\ x) pow\<^sub>\ y = a pow\<^sub>\ (x * y)" +proof - + have "1 \ inverse a" using assms using real_inverse_ge_one_lemma by simp + then have "(inverse a pow\<^sub>\ x) pow\<^sub>\ y = inverse a pow\<^sub>\ (x * y)" + using powreal_mult1 assms(3) by blast + then have "inverse((inverse a pow\<^sub>\ x) pow\<^sub>\ y) = a pow\<^sub>\ (x * y)" + by (simp add: assms(1) powreal_inverse) + then show ?thesis + using assms(1) powreal_gt_zero powreal_inverse by auto +qed + +lemma powreal_mult3: + "\ 0 < a; 0 \ x \ \ (a pow\<^sub>\ x) pow\<^sub>\ y = a pow\<^sub>\ (x * y)" +by (metis linorder_not_less powreal_mult1 powreal_mult2) + +lemma powreal_mult4: + assumes a0: "0 < a" and x0: "x \ 0" + shows "(a pow\<^sub>\ x) pow\<^sub>\ y = a pow\<^sub>\ (x * y)" +proof - + have minux0: "-x \ 0" using x0 by simp + then have "(a pow\<^sub>\ -x) pow\<^sub>\ y = a pow\<^sub>\ (-x * y)" + using powreal_mult3 a0 by simp + then have "inverse (a pow\<^sub>\ x) pow\<^sub>\ y = inverse (a pow\<^sub>\ (x * y))" + by (simp add: powreal_minus) + then have "inverse ((a pow\<^sub>\ x) pow\<^sub>\ y) = inverse (a pow\<^sub>\ (x * y))" + by (simp add: a0 powreal_gt_zero powreal_inverse) + then show ?thesis + using inverse_eq_iff_eq by blast +qed + +(* At long last... *) +lemma powreal_mult: + "(a pow\<^sub>\ x) pow\<^sub>\ y = a pow\<^sub>\ (x * y)" + by (smt (verit, best) powreal_def powreal_mult3 powreal_mult4) + +lemma powreal_divide: + "\ 0 \ a; 0 \ b \ \ (a/b) pow\<^sub>\ x = (a pow\<^sub>\ x) / (b pow\<^sub>\ x)" + by (simp add: divide_inverse powreal_inverse powreal_mult_base) + +lemma powreal_divide2: + "0 \ a \ a pow\<^sub>\ (x - y) = (a pow\<^sub>\ x) / (a pow\<^sub>\ y)" +proof - + assume a1: "0 \ a" + have f2: "\x0. - (x0::real) = - 1 * x0" + by auto + have "x - y = x + - 1 * y" + by auto + then show ?thesis + using f2 a1 by (metis field_class.field_divide_inverse powreal_add powreal_minus) +qed + +lemma powreal_less_mono_base: + assumes r0: "r > 0" and a0: "0 < a" and ab: "a < b" + shows "a pow\<^sub>\ r < b pow\<^sub>\ r" +proof - + have "b/a > 1" using ab a0 by simp + then have "(b/a) pow\<^sub>\ r > 1" + using powreal_gt_one r0 by simp + then have "b pow\<^sub>\ r / a pow\<^sub>\ r > 1" + using ab a0 powreal_divide by simp + also have "a pow\<^sub>\ r > 0" + by (simp add: a0 powreal_gt_zero) + ultimately show ?thesis + using less_divide_eq_1_pos by blast +qed + +lemma powreal_less_antimono_base: + assumes "r < 0" and "0 < a" and "a < b" + shows "a pow\<^sub>\ r > b pow\<^sub>\ r" +proof - + have "0 < -r" using assms(1) by simp + then have "a pow\<^sub>\ - r < b pow\<^sub>\ - r" + using assms(2) assms(3) powreal_less_mono_base by blast + then show ?thesis + using assms(2) assms(3) powreal_gt_zero powreal_minus by auto +qed + +lemma powa_power_eq: + assumes "a \ 1" shows "a powa (of_nat n) = a ^ n" +proof - + have "incseq (\m. rat_of_nat n)" by simp + moreover have "(\na. real_of_rat (rat_of_nat n)) \ real n" by simp + moreover have "(\na. a pow\<^sub>\ rat_of_nat n) \ a ^ n" + using powrat_power_eq assms by auto + ultimately have "(\na. a pow\<^sub>\ incseq_of (real n) na) \ a ^ n" + using powa_indep_incseq_of' assms by blast + then show ?thesis + by (metis Lim_def limI powa_def) +qed + +lemma powreal_power_eq: + "a > 0 \ a pow\<^sub>\ (of_nat n) = a ^ n" +unfolding powreal_def +by (simp add: powa_minus powa_power_eq power_inverse real_inverse_ge_one_lemma) + +lemma powreal_power_eq2: + "0 \ a \ 0 < n \ a ^ n = (if a = 0 then 0 else a pow\<^sub>\ (real n))" + by (auto simp add: powreal_power_eq) + +lemma powreal_mult_power: "a > 0 \ a pow\<^sub>\ (n * x) = (a pow\<^sub>\ x) ^ n" + by (metis mult.commute powreal_gt_zero powreal_mult powreal_power_eq) + +lemma powreal_int: + assumes "x > 0" + shows "x pow\<^sub>\ i = (if i \ 0 then x ^ nat i else 1 / x ^ nat (-i))" +proof cases + assume "i < 0" + have r: "x pow\<^sub>\ i = 1 / x pow\<^sub>\ (-i)" by (simp add: assms powreal_minus divide_inverse) + show ?thesis using `i < 0` `x > 0` by (simp add: r field_simps powreal_power_eq [symmetric]) +qed (simp add: assms powreal_power_eq[symmetric]) + +lemma powreal_numeral: "0 \ x \ x pow\<^sub>\ numeral n = x^numeral n" + using powreal_power_eq [of x "numeral n"] powreal_def + by fastforce + +lemma root_powreal_inverse: + assumes "0 < n" and "0 \ x" + shows "root n x = x pow\<^sub>\ (1/n)" +proof - + have "root n x = x pow\<^sub>\ real_of_rat (inverse (rat_of_nat n))" + using assms real_root_eq_powrat_inverse realpow_powrat_eq [symmetric] powreal_def + by simp + then show ?thesis + by (metis inverse_eq_divide of_rat_inverse of_rat_of_nat_eq) +qed + +lemma powreal_inverse_of_nat_gt_one: + "\ 1 < a; n \ 0\ \ a pow\<^sub>\ (inverse (of_nat n)) > 1" +by (metis inverse_positive_iff_positive neq0_conv of_nat_0_less_iff powreal_gt_one) + +end \ No newline at end of file diff --git a/thys/Real_Power/document/root.tex b/thys/Real_Power/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Real_Power/document/root.tex @@ -0,0 +1,40 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amsfonts, amsmath, amssymb} + +% this should be the last package used +\usepackage{pdfsetup} + +\begin{document} + +\title{Real Exponents as the Limits of Sequences of Rational Exponents} +\author{Jacques D. Fleuriot\\ + University of Edinburgh} +%\affil[1]{University of Edinburgh} +\maketitle + +\begin{abstract} +In this formalisation, we construct real exponents as the limits of sequences of rational exponents. In particular, if $a \ge 1$ and $x \in \mathbb{R}$, we choose an increasing rational sequence $r_n$ such that $\lim_{n\to\infty} {r_n} = x$. Then the sequence $a^{r_n}$ is increasing and if $r$ is any rational number such that $r > x$, $a^{r_n}$ is bounded above by $a^r$. By the convergence criterion for monotone sequences, $a^{r_n}$ converges. We define $a^ x = \lim_{n\to\infty} a^{r_n}$ and show that it has the expected properties (for $a \ge 0$). + +This particular construction of real exponents is needed instead of the usual one using the natural logarithm and exponential functions (which already exists in Isabelle) to support our mechanical derivation of Euler's exponential series as an ``infinite polynomial".\ Aside from helping us avoid circular reasoning, this is, as far as we are aware, the first time real exponents are mechanised in this way within a proof assistant. + +\end{abstract} + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +%generated text of all theories +\input{session} + +% optional bibliography +%\bibliographystyle{abbrv} +%\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/Szemeredi_Regularity/ROOT b/thys/Szemeredi_Regularity/ROOT new file mode 100644 --- /dev/null +++ b/thys/Szemeredi_Regularity/ROOT @@ -0,0 +1,12 @@ +chapter AFP + +session Szemeredi_Regularity (AFP) = "HOL-Library" + + options [timeout = 600] + sessions + "Girth_Chromatic" + theories + Szemeredi + document_files + "root.tex" + "root.bib" + diff --git a/thys/Szemeredi_Regularity/Szemeredi.thy b/thys/Szemeredi_Regularity/Szemeredi.thy new file mode 100644 --- /dev/null +++ b/thys/Szemeredi_Regularity/Szemeredi.thy @@ -0,0 +1,1185 @@ +section \Szemerédi's Regularity Lemma\ + +theory Szemeredi + imports Complex_Main "HOL-Library.Disjoint_Sets" "Girth_Chromatic.Ugraphs" + +begin + +text\We formalise Szemerédi's Regularity Lemma, which is a major result in the study of large graphs +(extremal graph theory). +We follow Yufei Zhao's notes ``Graph Theory and Additive Combinatorics'' (MIT) +\<^url>\https://ocw.mit.edu/courses/mathematics/18-217-graph-theory-and-additive-combinatorics-fall-2019/lecture-notes/MIT18_217F19_ch3.pdf\ +and W.T. Gowers's notes ``Topics in Combinatorics'' (University of Cambridge, Lent 2004, Chapter 3) +\<^url>\https://www.dpmms.cam.ac.uk/~par31/notes/tic.pdf\. +We also refer to a third source, also by Zhao and also entitled +``Graph Theory and Additive Combinatorics'': \<^url>\https://yufeizhao.com/gtac/gtac17.pdf\.\ + + +subsection \Miscellaneous\ + +text \A technical lemma used below in @{text \energy_graph_partition_half\}\ +lemma sum_products_le: + fixes a :: "'a \ 'b::linordered_idom" + assumes "\i. i \ I \ a i \ 0" + shows "(\i\I. a i * b i)\<^sup>2 \ (\i\I. a i) * (\i\I. a i * (b i)\<^sup>2)" + using assms +proof (induction I rule: infinite_finite_induct) + case (insert j I) + then have "(\i\insert j I. a i * b i)\<^sup>2 + \ (a j * b j)\<^sup>2 + 2 * a j * b j * (\i\I. a i * b i) + (\i\I. a i) * (\i\I. a i * (b i)\<^sup>2)" + using insert by (simp add: algebra_simps power2_eq_square) + also have "\ \ a j * (a j * (b j)\<^sup>2) + (a j * (sum a I * (b j)\<^sup>2 + (\i\I. a i * (b i)\<^sup>2)) + + sum a I * (\i\I. a i * (b i)\<^sup>2))" + proof - + have "0 \ (\i\I. a i * (b i - b j)\<^sup>2)" + by (simp add: insert.prems sum_nonneg) + then have "2 * b j * (\i\I. a i * b i) \ (sum a I * (b j)\<^sup>2) + (\i\I. a i * (b i)\<^sup>2)" + by (simp add: power2_eq_square algebra_simps sum_subtractf sum.distrib sum_distrib_left) + then show ?thesis + by (simp add: insert.prems power2_eq_square mult.commute mult.left_commute mult_left_mono) + qed + finally show ?case + by (simp add: insert algebra_simps) +qed auto + +subsubsection \Partitions indexed by integers\ + +definition finite_graph_partition :: "[uvert set, nat \ uvert set, nat] \ bool" + where "finite_graph_partition V P n \ partition_on V (P ` {.. disjoint_family_on P {..Equivalent, simpler characterisation\ +lemma finite_graph_partition_eq: "finite_graph_partition V P n \ partition_on V (P ` {.. inj_on P {.. disjoint_family_on P {.. V={}" + by (auto simp: finite_graph_partition_eq partition_on_def) + +lemma finite_graph_partition_equals: + "finite_graph_partition V P n \ (\ifinite_graph_partition V P n; i \ P i \ V" + using finite_graph_partition_equals by blast + +lemma finite_graph_partition_nonempty: + "\finite_graph_partition V P n; i \ P i \ {}" + by (metis finite_graph_partition_def image_eqI lessThan_iff partition_on_def) + +lemma trivial_graph_partition: + assumes "finite_graph_partition V P 1" shows "P 0 = V" + using finite_graph_partition_equals [OF assms] by auto + +lemma trivial_graph_partition_exists: + assumes "V \ {}" + shows "finite_graph_partition V (\i. V) (Suc 0)" + using assms by (simp add: lessThan_Suc finite_graph_partition_eq partition_on_def) + +lemma finite_graph_partition_finite: + assumes "finite_graph_partition V P k" "finite V" "i card V" + by (metis (mono_tags, lifting) UN_I assms card_mono finite_graph_partition_equals lessThan_iff subsetI) + +lemma finite_graph_partition_gt0: + assumes "finite_graph_partition V P k" "finite V" "i 0" + by (metis assms card_gt_0_iff finite_graph_partition_def finite_graph_partition_finite imageI lessThan_iff partition_on_def) + +lemma card_finite_graph_partition: + assumes "finite_graph_partition V P k" "finite V" + shows "(\ii = card (\ (P ` {.. = card V" + using assms finite_graph_partition_equals by blast + finally show ?thesis . +qed + +lemma finite_graph_partition_obtain: + assumes "finite_graph_partition V P k" "x \ V" + obtains i where "i < k" and "x \ P i" + using assms finite_graph_partition_equals by force + +text \Partitions into two parts\ +fun binary_partition :: "['a set, 'a set, nat] \ 'a set" + where "binary_partition A B 0 = A" + | "binary_partition A B (Suc 0) = B-A" + +lemma binary_partition2_eq [simp]: "binary_partition A B ` {..<2} = {A,B-A}" + by (auto simp add: numeral_2_eq_2 lessThan_Suc) + +lemma inj_binary_partition: + "B \ {} \ inj_on (binary_partition A B) {..<2}" + by (auto simp: numeral_2_eq_2 lessThan_Suc) + +lemma disjoint_family_binary_partition: "disjoint_family_on (binary_partition A B) {..<2}" + by (auto simp: numeral_2_eq_2 lessThan_Suc disjoint_family_on_def) + +lemma finite_graph_partition_binary_partition: + assumes "A \ B" "A \ {}" + shows "finite_graph_partition B (binary_partition A B) 2" + unfolding finite_graph_partition_def partition_on_def binary_partition2_eq + using assms + by (force simp add: disjnt_iff pairwise_insert disjoint_family_binary_partition) + +subsubsection \Tools to combine the refinements of the partition @{term "P i"} for each @{term i}\ + +text \These are needed to retain the ``intuitive'' idea of partitions as indexed by integers.\ + +lemma less_sum_nat_iff: + fixes b::nat and k::nat + shows "b < (\i (\ji b \ b < (\ii (\!j. j (\i b \ b < (\i nat) \ nat \ nat \ nat" + where "part a k b \ (THE j. j (\i b \ b < (\ii (let j = part a k b in j < k \ (\i < j. a i) \ b \ b < (\i < j. a i) + a j)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + unfolding less_sum_nat_iff_uniq part_def by (metis (no_types, lifting) theI') +qed (auto simp: less_sum_nat_iff Let_def) + +lemma part_Suc: "part a (Suc k) b = (if sum a {.. b \ b < sum a {..: "\j. \sum a {.. sum a {.. \ j = i" + by (meson \d < a i\ leD le_add1 less_sum_nat_iff nat_add_left_cancel_less not_less_iff_gr_or_eq) + show ?thesis + unfolding part_def + using assms + by (intro the_equality) (auto simp: \
) +qed + +lemma sum_layers_less: + fixes d::nat and k::nat + shows "\i < k; d < a i\ \ sum a {..Edges\ + +text \All edges between two sets of vertices, @{term X} and @{term Y}, in a graph, @{term G}\ + +definition all_edges_between :: "nat set \ nat set \ nat set \ nat set set \ (nat \ nat) set" + where "all_edges_between X Y G \ {(x,y). x\X \ y\Y \ {x,y} \ uedges G}" + +lemma all_edges_between_subset: "all_edges_between X Y G \ X\Y" + by (auto simp: all_edges_between_def) + +lemma max_all_edges_between: + assumes "finite X" "finite Y" + shows "card (all_edges_between X Y G) \ card X * card Y" + by (metis assms card_mono finite_SigmaI all_edges_between_subset card_cartesian_product) + +lemma all_edges_between_empty [simp]: + "all_edges_between {} Z G = {}" "all_edges_between Z {} G = {}" + by (auto simp: all_edges_between_def) + +lemma all_edges_between_disjnt1: + assumes "disjnt X Y" + shows "disjnt (all_edges_between X Z G) (all_edges_between Y Z G)" + using assms by (auto simp: all_edges_between_def disjnt_iff) + +lemma all_edges_between_disjnt2: + assumes "disjnt Y Z" + shows "disjnt (all_edges_between X Y G) (all_edges_between X Z G)" + using assms by (auto simp: all_edges_between_def disjnt_iff) + +lemma all_edges_between_Un1: + "all_edges_between (X \ Y) Z G = all_edges_between X Z G \ all_edges_between Y Z G" + by (auto simp: all_edges_between_def) + +lemma all_edges_between_Un2: + "all_edges_between X (Y \ Z) G = all_edges_between X Y G \ all_edges_between X Z G" + by (auto simp: all_edges_between_def) + +lemma finite_all_edges_between: + assumes "finite X" "finite Y" + shows "finite (all_edges_between X Y G)" + by (meson all_edges_between_subset assms finite_cartesian_product finite_subset) + +subsection \Edge Density and Regular Pairs\ + +text \The edge density between two sets of vertices, @{term X} and @{term Y}, in @{term G}. + Authors disagree on whether the sets are assumed to be disjoint!. + Quite a few authors assume disjointness, e.g. Malliaris and Shelah \<^url>\https://www.jstor.org/stable/23813167\ + For the following definitions, see pages 49--50 in Zhao's notes.\ +definition "edge_density X Y G \ card(all_edges_between X Y G) / (card X * card Y)" + +lemma edge_density_ge0: "edge_density X Y G \ 0" + by (auto simp: edge_density_def) + +lemma edge_density_le1: + assumes "finite X" "finite Y" + shows "edge_density X Y G \ 1" + using of_nat_mono [OF max_all_edges_between [OF assms]] + by (fastforce simp add: edge_density_def divide_simps) + +lemma all_edges_between_swap: + "all_edges_between X Y G = (\(x,y). (y,x)) ` (all_edges_between Y X G)" + unfolding all_edges_between_def + by (auto simp add: insert_commute image_iff split: prod.split) + +lemma card_all_edges_between_commute: + "card (all_edges_between X Y G) = card (all_edges_between Y X G)" +proof - + have "inj_on (\(x, y). (y, x)) A" for A :: "(nat*nat)set" + by (auto simp: inj_on_def) + then show ?thesis + by (simp add: all_edges_between_swap [of X Y] card_image) +qed + +lemma edge_density_commute: "edge_density X Y G = edge_density Y X G" + by (simp add: edge_density_def card_all_edges_between_commute mult.commute) + + +text \$\epsilon$-regular pairs, for two sets of vertices. Again, authors disagree on whether the +sets need to be disjoint, though it seems that overlapping sets cause double-counting. Authors also +disagree about whether or not to use the strict subset relation here. The proofs below are easier if +it is strict but later proofs require the non-strict version. The two definitions can be proved to +be equivalent under fairly mild conditions.\ + +definition regular_pair:: "uvert set \ uvert set \ ugraph \ real \ bool" + where "regular_pair X Y G \ \ + \A B. A \ X \ B \ Y \ (card A \ \ * card X) \ (card B \ \ * card Y) \ + \edge_density A B G - edge_density X Y G\ \ \" for \::real + +lemma regular_pair_commute: "regular_pair X Y G \ \ regular_pair Y X G \" + by (metis edge_density_commute regular_pair_def) + +abbreviation "irregular_pair X Y G \ \ \ regular_pair X Y G \" + +lemma irregular_pair_E: + fixes \::real + assumes "irregular_pair X Y G \" + obtains A B where "A \ X \ B \ Y \ (card A \ \ * card X) \ (card B \ \ * card Y)" + "\edge_density A B G - edge_density X Y G\ > \" + using assms by (auto simp: not_le regular_pair_def) + +lemma irregular_pair_I: + fixes \::real + assumes "A \ X" "B \ Y" "(card A \ \ * card X)" "(card B \ \ * card Y)" + "\edge_density A B G - edge_density X Y G\ > \" + shows "irregular_pair X Y G \" + using assms by (auto simp: not_le regular_pair_def) + +lemma edge_density_Un: + assumes "disjnt X1 X2" "finite X1" "finite X2" "finite Y" + shows "edge_density (X1 \ X2) Y G = (edge_density X1 Y G * card X1 + edge_density X2 Y G * card X2) / (card X1 + card X2)" + using assms unfolding edge_density_def + by (simp add: all_edges_between_disjnt1 all_edges_between_Un1 finite_all_edges_between card_Un_disjnt divide_simps) + +lemma edge_density_partition: + assumes U: "finite_graph_partition U P n" "finite U" and "finite W" + shows "edge_density U W G = (\ii (P ` {.. (P ` {.. (P ` {..ii (\i \ (P ` {.. = (edge_density (P n) W G * card (P n) + + edge_density (\ (P ` {.. (P ` {.. (P ` {.. (P ` {..finite W\ finU in auto) + also have "\ = (edge_density (P n) W G * card (P n) + + edge_density (\ (P ` {..ii = (\iiLet @{term P}, @{term Q} be partitions of a set of vertices @{term V}. + Then @{term P} refines @{term Q} if for all @{term \A \ P\} there is @{term \B \ Q\} + such that @{term \A \ B\}.\ + +definition finite_graph_refines :: "[uvert set, nat \ uvert set, nat, nat \ uvert set, nat] \ bool" + where "finite_graph_refines V P n Q m \ + finite_graph_partition V P n \ finite_graph_partition V Q m \ (\ij Q j)" + +text \For the sake of generality, and following Zhao's Online Lecture +\<^url>\https://www.youtube.com/watch?v=vcsxCFSLyP8&t=16s\ +we do not impose disjointness: we do not include @{term "i\j"} below.\ + +definition irregular_set:: "[real, ugraph, nat \ uvert set, nat] \ (nat\nat) set" + where "irregular_set \ \\::real. \G P k. {(i,j)|i j. i j irregular_pair (P i) (P j) G \}" + +text\A regular partition may contain a few irregular pairs as long as their total size is bounded as follows.\ +definition regular_partition:: "[real, ugraph, nat \ uvert set, nat] \ bool" + where + "regular_partition \ \\::real. \G P k. + finite_graph_partition (uverts G) P k \ + (\(i,j) \ irregular_set \ G P k. card (P i) * card (P j)) \ \ * (card (uverts G))\<^sup>2" + +lemma irregular_set_subset: "irregular_set \ G P k \ {.. {.. irregular_set \ G P k \ (j,i) \ irregular_set \ G P k" + by (auto simp add: irregular_set_def regular_pair_commute) + +lemma finite_graph_refines_refines: + assumes "finite_graph_refines V Y n X m" + shows "refines V (Y ` {.. G P k)" +proof - + have "irregular_set \ G P k \ {0.. {0..Energy of a Graph\ + +text \Definition 3.7 (Energy), written @{term "q(\,\)"}\ +definition energy_graph_subsets:: "[uvert set, uvert set, ugraph] \ real" where + "energy_graph_subsets \ \ G \ + card \ * card \ * (edge_density \ \ G)\<^sup>2 / (card (uverts G))\<^sup>2" + +text \Definition for partitions\ +definition energy_graph_partitions_subsets + :: "[ugraph, nat \ uvert set, nat, nat \ uvert set, nat] \ real" + where "energy_graph_partitions_subsets G U k W l \ + \ij \ G \ 0" + by (auto simp: energy_graph_subsets_def) + +lemma energy_graph_partitions_subsets_ge0 [simp]: + "energy_graph_partitions_subsets G U k W l \ 0" + by (auto simp: sum_nonneg energy_graph_partitions_subsets_def) + +lemma energy_graph_subsets_commute: + "energy_graph_subsets \ \ G = energy_graph_subsets \ \ G" + by (simp add: energy_graph_subsets_def edge_density_commute) + +lemma energy_graph_partitions_subsets_commute: + "energy_graph_partitions_subsets G W l U k = energy_graph_partitions_subsets G U k W l" + by (simp add: energy_graph_partitions_subsets_def energy_graph_subsets_commute sum.swap [where A="{..Definition 3.7 (Energy of a Partition), or following Gowers, mean square density: + a version of energy for a single partition of the vertex set. \ + +abbreviation mean_square_density :: "[ugraph, nat \ uvert set, nat] \ real" + where "mean_square_density G U k \ energy_graph_partitions_subsets G U k U k" + +lemma mean_square_density: + "mean_square_density G U k \ + (\ij2) + / (card (uverts G))\<^sup>2" + by (simp add: energy_graph_partitions_subsets_def energy_graph_subsets_def sum_divide_distrib) + +text\Observation: the energy is between 0 and 1 because the edge density is bounded above by 1.\ + +lemma sum_partition_le: + assumes "finite_graph_partition V P k" "finite V" + shows "(\ij (real(card V))\<^sup>2" + using assms +proof (induction k arbitrary: V) + case (Suc k) + then have [simp]: "V \ P k = P k" + using finite_graph_partition_equals by fastforce + have [simp]: "finite (P i)" if "i\k" for i + using Suc.prems finite_graph_partition_finite that by force + have C: "card (P i) \ card V" if "i\k" for i + using Suc.prems finite_graph_partition_le le_imp_less_Suc that by presburger + have D [simp]: "(\j (P ` {..ij (real (card (V - P k)))\<^sup>2" + by (simp add: finite_graph_partition_def lessThan_Suc partition_on_insert disjoint_family_on_insert comm_monoid_add_class.sum.distrib) + have "(\ijij \ real (card (P k) * card (P k)) + + 2 * (card V - card (P k)) * card (P k) + (real (card (V - P k)))\<^sup>2" + using * by linarith + also have "\ \ (real (card V))\<^sup>2" + by (simp add: of_nat_diff C card_Diff_subset_Int algebra_simps power2_eq_square) + finally show ?case . +qed auto + +lemma mean_square_density_bounded: + assumes "finite_graph_partition (uverts G) P k" "finite (uverts G)" + shows "mean_square_density G P k \ 1" +proof- + have \
: "edge_density (P i) (P j) G \ 1" for i j + using assms of_nat_mono [OF max_all_edges_between] + by (smt (verit, best) card.infinite divide_eq_0_iff edge_density_def edge_density_le1 mult_eq_0_iff of_nat_0) + have "(\ij2) + \ (\ij) + also have "\ \ (real(card (uverts G)))\<^sup>2" + using sum_partition_le using assms by presburger + finally show ?thesis + by (simp add: mean_square_density divide_simps) +qed + +subsection \Partitioning and Energy\ + +text\Zhao's Lemma 3.8 and Gowers's remark after Lemma 11. + Further partitioning of subsets of the vertex set cannot make the energy decrease. + We follow Gowers's proof, which avoids the use of probability.\ + +lemma energy_graph_partition_half: + assumes fin: "finite \" "finite \" and U: "finite_graph_partition \ U n" + shows "card \ * (edge_density \ \ G)\<^sup>2 \ (\i G)\<^sup>2)" +proof - + have "card \ * (edge_density \ \ G)\<^sup>2 = (\i \ G)\<^sup>2)" + by (metis \finite \\ U sum_distrib_right card_finite_graph_partition of_nat_sum) + also have "\ = edge_density \ \ G * (\i \ G * card (U i))" + by (simp add: sum_distrib_left power2_eq_square mult_ac) + also have "\ = (\i G * real (card (U i))) * edge_density \ \ G" + proof - + have "edge_density \ \ G * (\n G * card (U n)) + = edge_density \ \ G * (edge_density \ \ G * (\n = (\i G) * edge_density \ \ G" + by (simp add: sum_distrib_left mult_ac) + also have "\ = (\i G)\<^sup>2 / (\i \ (\i G)\<^sup>2)" + apply (clarsimp simp add: mult_ac divide_simps simp flip: of_nat_sum) + by (simp add: sum_products_le) + finally show ?thesis . +qed + +proposition energy_graph_partition_increase: + assumes fin: "finite \" "finite \" + and U: "finite_graph_partition \ U k" + and V: "finite_graph_partition \ W l" + shows "energy_graph_partitions_subsets G U k W l \ energy_graph_subsets \ \ G" +proof - + have "(card \ * card \) * (edge_density \ \ G)\<^sup>2 = card \ * (card \ * (edge_density \ \ G)\<^sup>2)" + using fin by (simp add: mult_ac) + also have "\ \ card \ * (\i G)\<^sup>2)" + by (intro mult_left_mono energy_graph_partition_half fin) (use assms in auto) + also have "\ = (\i * (edge_density \ (U i) G)\<^sup>2))" + by (simp add: sum_distrib_left edge_density_commute mult_ac) + also have "\ \ (\ij2))" + proof (intro mult_left_mono energy_graph_partition_half sum_mono fin) + show "finite (U i)" if "i \ {.. \ (\ij2)" + by (simp add: sum_distrib_left edge_density_commute mult_ac) + finally + have "(card \ * card \) * (edge_density \ \ G)\<^sup>2 + \ (\ij2)" . + then show ?thesis + unfolding energy_graph_partitions_subsets_def energy_graph_subsets_def + by (simp add: divide_simps flip: sum_divide_distrib) +qed + +lemma partition_imp_finite_graph_partition: + "\partition_on A P; finite P\ \ finite_graph_partition A (from_nat_into P) (card P)" + by (metis bij_betw_def bij_betw_from_nat_into_finite finite_graph_partition_eq) + +text \The following is the fully general version of Gowers's Lemma 11 and Zhao's Lemma 3.9. +Further partitioning of subsets of the vertex set cannot make the energy decrease. +Note that @{term V} should be @{term "uverts G"} even though this more general version holds.\ + +lemma energy_graph_partitions_subsets_increase_half: + assumes ref: "finite_graph_refines V Y n X m" and "finite V" + and finU: "\i. i < k \ card (U i) > 0" + shows "energy_graph_partitions_subsets G X m U k \ energy_graph_partitions_subsets G Y n U k" + (is "?lhs \ ?rhs") +proof - + have "\F. partition_on (X i) F \ F = Y ` {j. j Y j \ X i}" if "ii. i partition_on (X i) (F i) \ F i = Y ` {j. j Y j \ X i}" + by fastforce + have finite_X: "finite (X i)" if "ifinite V\ finite_graph_partition_finite finite_graph_refines_def ref that by auto + have finite_F: "finite (F i)" if "i {}" if "iiU\F i. f U) = (\i real" + proof - + have Yn_eq: "Y ` {..ii. {Y i}) {..ii = (sum \ sum) f (F ` {.. = (\iU\F i. f U)" + unfolding comp_apply by (metis injF sum.reindex_cong) + finally show ?thesis + by simp + qed + have "?lhs = (\ij \ (\ijx. U j) 1)" + proof - + have "finite_graph_partition (X i) (from_nat_into (F i)) (card (F i))" + if "i < m" "j < k" for i j + using partition_imp_finite_graph_partition that F by fastforce + moreover have "finite_graph_partition (U j) (\x. U j) (Suc 0)" + if "i < m" "j < k" for i j + using that by (metis card.empty finU less_not_refl trivial_graph_partition_exists) + ultimately show ?thesis + by (intro sum_mono energy_graph_partition_increase; simp add: finU card_ge_0_finite finite_X) + qed + also have "\ = (\irj = (\iv\F i. \j = ?rhs" + by (simp add: energy_graph_partitions_subsets_def F_sums_Y) + finally show ?thesis . +qed + +proposition energy_graph_partitions_subsets_increase: + assumes refX: "finite_graph_refines V Y n X m" and refU: "finite_graph_refines V' W l U k" + and "finite V" "finite V'" + shows "energy_graph_partitions_subsets G X m U k \ energy_graph_partitions_subsets G Y n W l" + (is "?lhs \ ?rhs") +proof - + have finU: "\i. i < k \ card (U i) > 0" and finY: "\i. i < n \ card (Y i) > 0" + using finite_graph_partition_gt0 finite_graph_refines_def assms by auto + have "?lhs \ energy_graph_partitions_subsets G Y n U k" + using assms energy_graph_partitions_subsets_increase_half finU refX by blast + also have "\ = energy_graph_partitions_subsets G U k Y n" + by (meson energy_graph_partitions_subsets_commute) + also have "\ \ energy_graph_partitions_subsets G W l Y n" + using assms energy_graph_partitions_subsets_increase_half finY refU by blast + also have "\ = ?rhs" + by (simp add: energy_graph_partitions_subsets_commute) + finally show ?thesis . +qed + +text \The original version of Gowers's Lemma 11 and Zhao's Lemma 3.9 + is not general enough to be used for anything.\ +corollary mean_square_density_increase: + assumes "finite_graph_refines V Y n X m" "finite V" + shows "mean_square_density G X m \ mean_square_density G Y n" + using assms energy_graph_partitions_subsets_increase by presburger + + +text\The Energy Boost Lemma (Lemma 3.10 in Zhao's notes) says that an +irregular partition increases the energy substantially. We assume that @{term "\ \ uverts G"} +and @{term "\ \ uverts G"} are not irregular, as witnessed by their subsets @{term"U1 \ \"} and @{term"W1 \ \"}. +The proof follows Lemma 12 of Gowers. \ + +proposition energy_boost: + fixes \::real and \ \ G + defines "alpha \ edge_density \ \ G" + defines "u \ \X Y. edge_density X Y G - alpha" + assumes "finite \" "finite \" + and "U1 \ \" "W1 \ \" "\ > 0" + and U1: "card U1 \ \ * card \" and W1: "card W1 \ \ * card \" + and gt: "\u U1 W1\ > \" + shows "(\A \ {U1, \ - U1}. \B \ {W1, \ - W1}. energy_graph_subsets A B G) + \ energy_graph_subsets \ \ G + \^4 * (card \ * card \) / (card (uverts G))\<^sup>2" + (is "?lhs \ ?rhs") +proof - + define UF where "UF \ {U1, \ - U1}" + define WF where "WF \ {W1, \ - W1}" + obtain [simp]: "finite \" "finite \" + using assms by (meson finite_subset) + then obtain 0: "card \ > 0" "card \ > 0" + using assms by fastforce + then obtain 1: "card U1 > 0" "card W1 > 0" + by (smt (verit) U1 W1 \\ > 0\ of_nat_0_less_iff zero_less_mult_iff) + then obtain [simp]: "finite U1" "finite W1" + by (meson card_ge_0_finite) + have 2 [simp]: "card x > 0" if "x \ UF" for x + using "1"(1) UF_def assms that by fastforce + have 3 [simp]: "card x > 0" if "x \ WF" for x + using "1"(2) WF_def assms that by force + have cardUW: "card \ = card U1 + card(\ - U1)" "card \ = card W1 + card(\ - W1)" + using 0 1 \U1 \ \\ \W1 \ \\ + by (metis card_eq_0_iff card_Diff_subset card_mono le_add_diff_inverse less_le)+ + + have CU: "card (all_edges_between \ Z G) + = card (all_edges_between (\ - U1) Z G) + card (all_edges_between U1 Z G)" + if "finite Z" for Z + using card_Un_disjnt all_edges_between_Un1 all_edges_between_disjnt1 \U1 \ \\ that + by (metis DiffD2 Un_Diff_cancel2 \finite U1\ \finite \\ disjnt_iff finite_Diff + finite_all_edges_between sup.strict_order_iff) + + have CW: "card (all_edges_between Z \ G) + = card (all_edges_between Z (\ - W1) G) + card (all_edges_between Z W1 G)" + if "finite Z" for Z + using card_Un_disjnt all_edges_between_Un2 all_edges_between_disjnt2 \W1 \ \\ that + by (metis DiffD2 Un_Diff_cancel2 \finite W1\ \finite \\ disjnt_iff finite_Diff + finite_all_edges_between sup.strict_order_iff) + have [simp]: "U1 \ \ - U1" "W1 \ \ - W1" + using assms by blast+ + have *: "(\i\UF. \j\WF. real (card (all_edges_between i j G))) + = card (all_edges_between \ \ G)" + by (simp add: UF_def WF_def cardUW CU CW) + have **: "real (card \) * real (card \) = (\i\UF. \j\WF. card i * card j)" + by (simp add: UF_def WF_def cardUW algebra_simps) + + let ?S = "\i\UF. \j\WF. (card i * card j) / (card \ * card \) * (edge_density i j G)\<^sup>2" + have \
: "2 * (\i\UF. \j\WF. + (card i * card j) / (card \ * card \) * (edge_density i j G)) + = alpha + alpha * (\i\UF. \j\WF. (card i * card j) / (card \ * card \))" + unfolding alpha_def + by (simp add: * ** edge_density_def divide_simps flip: sum_divide_distrib) + + have "\ * \ \ u U1 W1 * u U1 W1" + by (metis abs_ge_zero abs_mult_self_eq \\ > 0\ gt less_le mult_mono) + then have "(\*\)*(\*\) \ (card U1 * card W1) / (card \ * card \) * (u U1 W1)\<^sup>2" + using 0 mult_mono [OF U1 W1] \\ > 0\ + apply (simp add: divide_simps eval_nat_numeral) + by (smt (verit, del_insts) mult.assoc mult.commute mult_mono' of_nat_0_le_iff zero_le_mult_iff) + also have "\ \ (\i\UF. \j\WF. (card i * card j) / (card \ * card \) * (u i j)\<^sup>2)" + by (simp add: UF_def WF_def) + also have "\ = ?S - 2 * alpha * (\i\UF. \j\WF. + (card i * card j) / (card \ * card \) * edge_density i j G) + + alpha\<^sup>2 * (\i\UF. \j\WF. (card i * card j) / (card \ * card \))" + by (simp add: u_def power2_diff mult_ac ring_distribs divide_simps + sum_distrib_left sum_distrib_right sum_subtractf sum.distrib flip: sum_divide_distrib) + also have "\ = ?S - alpha\<^sup>2" + using \
by (simp add: power2_eq_square algebra_simps) + finally have 12: "alpha\<^sup>2 + \^4 \ ?S" + by (simp add: eval_nat_numeral) + have "?rhs = (alpha\<^sup>2 + \^4) * (card \ * card \ / (card (uverts G))\<^sup>2)" + unfolding alpha_def energy_graph_subsets_def + by (simp add: ring_distribs divide_simps power2_eq_square) + also have "\ \ ?S * (card \ * card \ / (card (uverts G))\<^sup>2)" + by (rule mult_right_mono [OF 12]) auto + also have "\ = ?lhs" + using 0 unfolding energy_graph_subsets_def UF_def WF_def + by (auto simp add: algebra_simps) + finally show ?thesis . +qed + +subsection \Towards Zhao's Lemma 3.11\ + +text\Lemma 3.11 says that we can always find a refinement +that increases the energy by a certain amount.\ + +text \A necessary lemma for the tower of exponentials in the result. Angeliki's proof\ +lemma le_tower_2: "k * (2 ^ Suc k) \ 2^(2^k)" +proof (induction k rule: less_induct) + case (less k) + show ?case + proof (cases "k \ Suc (Suc 0)") + case False + define j where "j = k - Suc 0" + have kj: "k = Suc j" + using False j_def by force + have "(3::nat) \ 2 ^ j" + by (metis kj False Suc_leI less_trans_Suc less_exp not_less numeral_3_eq_3) + then have \
: "(2^j + 3) \ (2::nat) ^ k" + by (simp add: kj) + have "k * (2 ^ Suc k) \ 6 * j * 2^j" + using False by (simp add: kj) + also have "\ \ 6 * 2^(2^j)" + using kj less.IH by force + also have "\ < 8 * 2^(2^j)" by simp + also have "\ = 2^(2^j + 3)" + by (simp add: power_add) + also have "\ \ 2^2^k" + using \
by (metis One_nat_def less_2_cases_iff power_increasing_iff) + finally show ?thesis + by simp + qed (auto simp: le_Suc_eq) +qed + + +text \Zhao's actual Lemma 3.11. However, the bound $m \le k 2 ^{k+1}$ +comes from a different source by Zhao: ``Graph Theory and Additive Combinatorics'', \<^url>\https://yufeizhao.com/gtac/gtac17.pdf\. +Zhao's original version, and Gowers', both have incorrect bounds.\ +proposition exists_refinement: + assumes "finite_graph_partition (uverts G) P k" and "finite (uverts G)" + and irreg: "\ regular_partition \ G P k" and "\ > 0" + obtains Q m where "finite_graph_refines (uverts G) Q m P k" + "mean_square_density G Q m \ mean_square_density G P k + \^5" + "\i. i card {j. j Q j \ P i} \ 2 ^ Suc k" + "m \ k * 2 ^ Suc k" +proof - + define sum_pp where "sum_pp \ (\(i,j) \ irregular_set \ G P k. card (P i) * card (P j))" + have "k \ 0" + using assms unfolding regular_partition_def irregular_set_def + by (intro notI) auto + then have G_nonempty: "0 < card (uverts G)" + by (metis assms(1) assms(2) finite_graph_partition_gt0 finite_graph_partition_le gr_zeroI leD) + have part_GP: "partition_on (uverts G) (P ` {.. {}" if "i \ * (card (uverts G))\<^sup>2" + using assms by (metis not_le regular_partition_def sum_pp_def) + then have sum_irreg_pos: "sum_pp > 0" + using \\ > 0\ G_nonempty less_asym by fastforce + have "\X\P i. + \Y\P j. + \ * card (P i) \ card X \ + \ * card (P j) \ card Y \ + \edge_density X Y G - edge_density (P i) (P j) G\ > \" + if "(i,j) \ irregular_set \ G P k" for i j + using that assms(1) finite_graph_partition_subset by (simp add: irregular_set_def regular_pair_def not_le) + then obtain X0 Y0 + where XY0_psub_P: "\i j. \(i,j) \ irregular_set \ G P k\ \ X0 i j \ P i \ Y0 i j \ P j" + and XY0_eps: + "\i j. \(i,j) \ irregular_set \ G P k\ + \ \ * card (P i) \ card (X0 i j) \ + \ * card (P j) \ card (Y0 i j) \ + \edge_density (X0 i j) (Y0 i j) G - edge_density (P i) (P j) G\ > \" + by metis + + define X where "X \ \i j. if j>i then Y0 j i else X0 i j" + define Y where "Y \ \i j. if j>i then X0 j i else Y0 i j" + + have XY_psub_P: "\i j. \(i,j) \ irregular_set \ G P k\ \ X i j \ P i \ Y i j \ P j" + using XY0_psub_P by (force simp: X_def Y_def irregular_set_swap) + have XY_eps: + "\i j. \(i,j) \ irregular_set \ G P k\ + \ \ * card (P i) \ card (X i j) \ + \ * card (P j) \ card (Y i j) \ + \edge_density (X i j) (Y i j) G - edge_density (P i) (P j) G\ > \" + using XY0_eps by (force simp: X_def Y_def edge_density_commute irregular_set_swap) + + have cardP: "card (P i) > 0" if "i {}" "Y i j \ {}" if "(i,j) \ irregular_set \ G P k" for i j + using XY_eps [OF that] that \\ > 0\ cardP [of i] cardP [of j] + by (auto simp: irregular_set_def mult_le_0_iff) + + text\By the assumption that our partition is irregular, there are many irregular pairs. + For each irregular pair, find pairs of subsets that witness irregularity.\ + define XP where "XP i \ ((\j. {X i j, P i - X i j}) ` {j. (i,j) \ irregular_set \ G P k})" for i + define YP where "YP j \ ((\i. {Y i j, P j - Y i j}) ` {i. (i,j) \ irregular_set \ G P k})" for j + + text \include degenerate partition to ensure it works whether or not there's an irregular pair\ + define PP where "PP \ \i. insert {P i} (XP i \ YP i)" + have finPP: "finite (PP i)" for i + unfolding PP_def XP_def YP_def + by (auto simp: inj_def intro!: finite_inverse_image finite_irregular_set) + have inPP_fin: "P \ PP i \ finite P" for P i + by (auto simp: PP_def XP_def YP_def) + have fin_cf: "finite (common_refinement (PP i))" for i + by (simp add: finPP finite_common_refinement inPP_fin) + + have part_cr: "partition_on (P i) (common_refinement (PP i))" if "i < k" for i + proof (intro partition_on_common_refinement partition_onI) + show "\\. \ \ PP i \ {} \ \" + using that XY_nonempty XY_psub_P finP(2) by (fastforce simp add: PP_def XP_def YP_def) + qed (auto simp: disjnt_iff PP_def XP_def YP_def dest: XY_psub_P) + + define QS where "QS i \ from_nat_into (common_refinement (PP i))" for i + define r where "r i \ card (common_refinement (PP i))" for i + have to_nat_on_cr: + "\x i. \x \ common_refinement (PP i); i < k\ + \ to_nat_on (common_refinement (PP i)) x < r i" + by (metis bij_betw_apply fin_cf lessThan_iff r_def to_nat_on_finite) + have QSr_eq: "QS i ` {.. {}" if "i P i" if "i < k" "q < r i" for i q + using part_cr [of i] QSr_eq that + unfolding partition_on_def by blast + then have QS_inject: "i = i'" if "QS i q = QS i' q'" "i \j. (\i sumr k" + define Q where "Q \ \d. QS (part r k d) (d - sumr (part r k d))" + have QS_Q: "QS j q = Q (sumr j + q)" if "ji ?rhs" + using QSr_eq r_def + apply (simp add: image_subset_iff Ball_def Bex_def QS_def Q_def m_def) + by (metis sumr_def add.right_neutral card.empty from_nat_into leD part) + show "?rhs \ ?lhs" + using QS_Q m_def sum_layers_less sumr_def by fastforce + qed + have Qm_eq: "Q ` {..i (Q ` {..i (common_refinement (PP i)))" + by (auto simp add: Qm_eq) + also have "\ = uverts G" + using part_cr finite_graph_partition_equals assms(1) + by (force simp: partition_on_def) + finally show "\ (Q ` {.. Q ` {.. Q ` {.. q" for p q + proof - + from that + obtain i j where "i common_refinement (PP i)" and j: "q \ common_refinement (PP j)" + by (auto simp add: Qm_eq) + show ?thesis + proof (cases "i=j") + case True + then show ?thesis + using part_cr [of i] + by (metis \j < k\ i j pairwise_def partition_on_def \p \ q\) + next + case False + have "p \ P i \ q \ P j" + by (metis Union_upper \i < k\ \j < k\ i j part_cr partition_onD1) + then show ?thesis + using False \i < k\ \j < k\ part_GP disjfam_P + by (metis disjnt_Union1 disjnt_Union2 disjnt_def disjoint_family_on_def i j lessThan_iff part_cr partition_onD1) + qed + qed + show "{} \ Q ` {..i < part r k i. r i) \ i \ i < (\i < part r k i. r i) + r (part r k i)" + by (metis \i < m\ m_def part sumr_def)+ + then have "QS (part r k i) (i - sum r {.. P (part r k i)" + by (simp add: QS_subset_P less_diff_conv2) + then show "\j P j" + unfolding Q_def sumr_def using \part r k i < k\ by blast + qed (use assms in auto) + + have inj_Q: "inj_on Q {.. (\(i,j) \ irregular_set \ G P k. \^4 * (card (P i) * card (P j)) / (card (uverts G))\<^sup>2)" + { fix i j + assume ij: "(i,j) \ irregular_set \ G P k" + then have "i^4 * (card (P i) * card (P j)) / (card (uverts G))\<^sup>2 + \ (\A \ {X i j, P i - X i j}. \B \ {Y i j, P j - Y i j}. energy_graph_subsets A B G)" + using XY_psub_P [OF ij] XY_eps [OF ij] assms + by (intro energy_boost \i < k\ \j < k\ finP \\>0\) auto + also have "\ = (\a<2. \b<2. energy_graph_subsets (binary_partition (X i j) (P i) a) (binary_partition (Y i j) (P j) b) G)" + using \i < k\ \j < k\ + by (simp add: sum.reindex inj_binary_partition finP flip: binary_partition2_eq) + also have "\ = energy_graph_partitions_subsets G (binary_partition (X i j) (P i)) 2 (binary_partition (Y i j) (P j)) 2" + by (simp add: energy_graph_partitions_subsets_def) + finally have "energy_graph_subsets (P i) (P j) G + \^4 * (card (P i) * card (P j)) / (card (uverts G))\<^sup>2 + \ energy_graph_partitions_subsets G (binary_partition (X i j) (P i)) 2 (binary_partition (Y i j) (P j)) 2" . + } note A = this + { fix i j + assume ij: "(i,j) \ irregular_set \ G P k" + then have "i PP i" + using ij by (simp add: PP_def XP_def) + have I: "finite_graph_partition (P i) (QS i) (r i)" + by (simp add: QSr_eq \i < k\ finite_graph_partition_eq inj_QS part_cr) + moreover have "\qb<2. QS i q \ binary_partition (X i j) (P i) b" + using common_refinement_exists [OF _ XPX] + by (fastforce simp add: QS_def r_def numeral_2_eq_2 intro: from_nat_into) + ultimately have ref_XP: "finite_graph_refines (P i) (QS i) (r i) (binary_partition (X i j) (P i)) 2" + unfolding finite_graph_refines_def + using XY_nonempty(1) XY_psub_P finite_graph_partition_binary_partition ij by presburger + have YPY: "{Y i j, P j - Y i j} \ PP j" + using ij by (simp add: PP_def YP_def) + have J: "finite_graph_partition (P j) (QS j) (r j)" + by (simp add: QSr_eq \j < k\ finite_graph_partition_eq inj_QS part_cr) + moreover have "\qb<2. QS j q \ binary_partition (Y i j) (P j) b" + using common_refinement_exists [OF _ YPY] + by (fastforce simp add: QS_def r_def numeral_2_eq_2 intro: from_nat_into) + ultimately have ref_YP: "finite_graph_refines (P j) (QS j) (r j) (binary_partition (Y i j) (P j)) 2" + unfolding finite_graph_refines_def + using XY_nonempty(2) XY_psub_P finite_graph_partition_binary_partition ij by presburger + have "energy_graph_partitions_subsets G (binary_partition (X i j) (P i)) 2 (binary_partition (Y i j) (P j)) 2 + \ energy_graph_partitions_subsets G (QS i) (r i) (QS j) (r j)" + using \i < k\ \j < k\ + by (simp add: finP energy_graph_partitions_subsets_increase [OF ref_XP ref_YP]) + } note B = this + have df_QS: "disjoint_family_on (\i. QS i ` {..^5 \ mean_square_density G P k + sum_eps" + proof - + have "\^5 = (\ * (card (uverts G))\<^sup>2) * (\^4 / (card (uverts G))\<^sup>2)" + using G_nonempty by (simp add: field_simps eval_nat_numeral) + also have "\ \ sum_pp * (sum_eps / sum_pp)" + proof (rule mult_mono) + show "\^4 / real ((card (uverts G))\<^sup>2) \ sum_eps / sum_pp" + using sum_irreg_pos sum_eps_def sum_pp_def + by (auto simp add: case_prod_unfold sum.neutral simp flip: sum_distrib_left sum_divide_distrib of_nat_sum of_nat_mult) + qed (use spp sum_nonneg in auto) + also have "\ \ sum_eps" + by (simp add: sum_irreg_pos) + finally show ?thesis by simp + qed + also have "\ = (\(i,j)\?REG. energy_graph_subsets (P i) (P j) G) + + (\(i,j)\irregular_set \ G P k. energy_graph_subsets (P i) (P j) G) + sum_eps" + by (simp add: energy_graph_partitions_subsets_def sum.cartesian_product irregular_set_subset sum.subset_diff) + also have "\ \ (\(i,j) \ ?REG. energy_graph_subsets (P i) (P j) G) + + (\(i,j) \ irregular_set \ G P k. energy_graph_partitions_subsets G (binary_partition (X i j) (P i)) 2 (binary_partition (Y i j) (P j)) 2)" + using A unfolding sum_eps_def case_prod_unfold + by (force intro: sum_mono simp flip: energy_graph_partitions_subsets_def sum.distrib) + also have "\ \ (\(i,j) \ ?REG. energy_graph_partitions_subsets G (QS i) (r i) (QS j) (r j)) + + (\(i,j) \ irregular_set \ G P k. energy_graph_partitions_subsets G (binary_partition (X i j) (P i)) 2 (binary_partition (Y i j) (P j)) 2)" + by(auto simp: part_P_QS finP intro!: sum_mono energy_graph_partition_increase) + also have "\ \ (\(i,j) \ ?REG. energy_graph_partitions_subsets G (QS i) (r i) (QS j) (r j)) + + (\(i,j) \ irregular_set \ G P k. energy_graph_partitions_subsets G (QS i) (r i) (QS j) (r j))" + using B + proof (intro sum_mono add_mono ordered_comm_monoid_add_class.sum_mono2) + qed (auto split: prod.split) + also have "\ = (\(i,j) \ ?KK. energy_graph_partitions_subsets G (QS i) (r i) (QS j) (r j))" + by (metis (no_types, lifting) finite_SigmaI finite_lessThan irregular_set_subset sum.subset_diff) + also have "\ = (\ij = (\A \ Q ` {..B \ Q ` {.. = (\ij = mean_square_density G Q m" + by (simp add: mean_square_density energy_graph_subsets_def sum_divide_distrib) + finally show "mean_square_density G P k + \ ^ 5 \ mean_square_density G Q m" . + + let ?QinP = "\i. {j. j < m \ Q j \ P i}" + show card_QP: "card (?QinP i) \ 2 ^ Suc k" + if "i < k" for i + proof - + have card_cr: "card (common_refinement (PP i)) \ 2 ^ Suc k" + proof - + have "card (common_refinement (PP i)) \ prod card (PP i)" + by (intro card_common_refinement finPP inPP_fin) + also have "\ = prod card (XP i \ YP i)" + unfolding PP_def + proof (subst prod.insert) + have XYP_2: "\i \. \ \ XP i \ YP i \ card \ = 2" + by (auto simp: XP_def YP_def; metis Diff_iff XY_psub_P card_2_iff psubset_imp_ex_mem) + then show "{P i} \ XP i \ YP i" + by fastforce + qed (use PP_def finPP in auto) + also have "\ \ 2 ^ Suc k" + proof (rule prod_le_power) + define XS where "XS \ (\l. {X0 i l, P i - X0 i l}) ` {..i}" + define YS where "YS \ (\l. {Y0 l i, P i - Y0 l i}) ` {i.. Suc i" "card YS \ k-i" + unfolding XS_def YS_def using card_image_le by force+ + with that have k': "card XS + card YS \ Suc k" + by linarith + have finXYS: "finite (XS \ YS)" + by (simp add: XS_def YS_def) + have "XP i \ YP i \ XS \ YS" + by (auto simp: XP_def YP_def X_def Y_def XS_def YS_def irregular_set_def) + then have "card (XP i \ YP i) \ card XS + card YS" + by (meson card_Un_le card_mono finXYS order_trans) + then show "card (XP i \ YP i) \ Suc k" + using k' le_trans by blast + fix x + assume "x \ XP i \ YP i" + then show "0 \ card x \ card x \ 2" + by (auto simp: XP_def YP_def card_insert_le_m1) + qed auto + finally show ?thesis . + qed + have "i' = i" if "QS i' q \ P i" "i'i < k\ inf.bounded_iff lessThan_iff subset_empty that) + then have "?QinP i \ {sumr i..< sumr (Suc i)}" + by (clarsimp simp: Q_def sumr_def m_def) (metis add_less_cancel_left le_add_diff_inverse part) + then have "card (?QinP i) \ card {sumr i..< sumr (Suc i)}" + by (meson card_mono finite_atLeastLessThan) + also have "\ \ 2 ^ Suc k" + using card_cr by (simp add: sumr_def r_def) + finally show ?thesis . + qed + have "m = card {j. j < m}" + by simp + also have "\ \ card (\i (\i \ (\i \ (\i k * 2 ^ Suc k" + by simp + qed +qed + +subsection \The Regularity Proof Itself\ + +text\Szemerédi's Regularity Lemma is Theorem 3.5 in Zhao's notes.\ + +text\We start with a trivial partition (one part). If it is already $\epsilon$-regular, we are done. If +not, we refine it by applying lemma @{thm[source]"exists_refinement"} above, which increases the +energy. We can repeat this step, but it cannot increase forever: by @{thm [source] +mean_square_density_bounded} it cannot exceed~1. This defines an algorithm that must stop +after at most $\epsilon^{-5}$ steps, resulting in an $\epsilon$-regular partition.\ +theorem Szemeredi_Regularity_Lemma: + assumes "\ > 0" + obtains M where "\G. card (uverts G) > 0 \ \P k. regular_partition \ G P k \ k \ M" +proof + fix G + assume "card (uverts G) > 0" + then obtain finG: "finite (uverts G)" and nonempty: "uverts G \ {}" + by (simp add: card_gt_0_iff) + define \ where "\ \ \Q m P k. finite_graph_refines (uverts G) Q m P k \ + mean_square_density G Q m \ mean_square_density G P k + \^5 \ + m \ k * 2 ^ Suc k" + define nxt where "nxt \ \(P,k). if regular_partition \ G P k then (P,k) else SOME (Q,m). \ Q m P k" + define iter where "iter \ \i. (nxt ^^ i) ((\u. uverts G, Suc 0))" + define last where "last \ Suc (nat\1 / \ ^ 5\)" + have iter_Suc [simp]: "iter (Suc i) = nxt (iter i)" for i + by (simp add: iter_def) + have \: "\ Q m P k" + if Pk: "finite_graph_partition (uverts G) P k" and irreg: "\ regular_partition \ G P k" + and "(Q,m) = nxt (P,k)" for P k Q m + using that exists_refinement [OF Pk finG irreg assms] irreg + apply (simp add: nxt_def \_def) + by (metis (mono_tags, lifting) case_prod_conv someI) + have iter_finite_graph_partition: "case iter i of (P,k) \ finite_graph_partition (uverts G) P k" for i + proof (induction i) + case 0 + then show ?case + by (simp add: iter_def nonempty trivial_graph_partition_exists) + next + case (Suc i) + with \ show ?case + apply (simp add: nxt_def \_def split: prod.split) + by (metis (no_types, lifting) case_prodD finite_graph_refines_def) + qed + have False if irreg: "\i. i\last \ (case iter i of (P,k) \ \ regular_partition \ G P k)" + proof - + have \_loop: "\ Q m P k" if "nxt (P,k) = (Q,m)" "iter i = (P,k)" "i\last" for i Q m P k + by (metis \ case_prod_conv irreg iter_finite_graph_partition that) + have iter_grow: "(case iter i of (Q,m) \ mean_square_density G Q m) \ i * \^5" if "i\last" for i + using that + proof (induction i) + case (Suc i) + then show ?case + by (auto simp add: algebra_simps \_def split: prod.split prod.split_asm dest: \_loop) + qed (auto simp: iter_def) + have "last * \^5 \ (case iter last of (Q,m) \ mean_square_density G Q m)" + by (simp add: iter_grow) + also have "\ \ 1" + by (metis (no_types, lifting) finG case_prod_beta iter_finite_graph_partition mean_square_density_bounded) + finally have "real last * \ ^ 5 \ 1" . + with assms show False + unfolding last_def by (meson lessI natceiling_lessD not_less pos_divide_less_eq zero_less_power) + qed + then obtain i where "i \ last" and "case iter i of (P,k) \ regular_partition \ G P k" + by force + then have reglar: "case iter (i + d) of (P,k) \ regular_partition \ G P k" for d + by (induction d) (auto simp add: nxt_def split: prod.split prod.split_asm) + define tower where "tower \ \k. (power(2::nat) ^^ k) 2" + have [simp]: "tower (Suc k) = 2 ^ tower k" for k + by (simp add: tower_def) + have iter_tower: "(case iter i of (Q,m) \ m) \ tower (2*i)" for i + proof (induction i) + case 0 + then show ?case + by (auto simp: iter_def tower_def) + next + case (Suc i) + then obtain Q m P k where Qm: "nxt (P,k) = (Q,m)" "iter i = (P,k)" and kle: "k \ tower (2 * i)" + by (metis case_prod_conv surj_pair) + then have *: "m \ k * 2 ^ Suc k" + using iter_finite_graph_partition [of i] \ + by (smt (verit, ccfv_SIG) Suc_leD \_def case_prod_conv le_eq_less_or_eq mult_le_mono2 nxt_def + power2_eq_square power2_nat_le_imp_le prod.simps(1) self_le_ge2_pow) + also have "\ \ 2 ^ 2 ^ tower (2 * i)" + by (metis (full_types) One_nat_def le_tower_2 lessI numeral_2_eq_2 order.trans power_increasing_iff kle) + finally show ?case + by (simp add: Qm) + qed + then show "\P k. regular_partition \ G P k \ k \ tower(2 * last)" + using reglar \i \ last\ by (metis case_prod_beta nat_le_iff_add) +qed + +text \The actual value of the bound is visible above: a tower of exponentials of height $2(1 + \epsilon^{-5})$.\ + +end diff --git a/thys/Szemeredi_Regularity/document/root.bib b/thys/Szemeredi_Regularity/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Szemeredi_Regularity/document/root.bib @@ -0,0 +1,47 @@ +%% This BibTeX bibliography file was created using BibDesk. +%% http://bibdesk.sourceforge.net/ + + +%% Created for Larry Paulson at 2021-11-04 21:29:28 +0000 + + +%% Saved with string encoding Unicode (UTF-8) + + + +@book{diestel-graph, + author = {Reinhard Diestel}, + booktitle = {Graph Theory}, + date-added = {2021-11-04 21:28:13 +0000}, + date-modified = {2021-11-04 21:29:19 +0000}, + publisher = {Springer}, + title = {Graph Theory}, + year = {2017}} + +@techreport{szemeredi-regular, + author = {Endre Szemer{\'e}di}, + date-added = {2021-11-04 20:50:28 +0000}, + date-modified = {2021-11-04 20:54:26 +0000}, + institution = {Stanford University Computer Science Department}, + month = apr, + number = {STAN-CS-75-489}, + title = {Regular Partitions of Graphs}, + url = {http://i.stanford.edu/pub/cstr/reports/cs/tr/75/489/CS-TR-75-489.pdf}, + year = {1975}, + bdsk-url-1 = {http://i.stanford.edu/pub/cstr/reports/cs/tr/75/489/CS-TR-75-489.pdf}} + +@article{malliaris-regularity, + abstract = {We develop a framework in which Szemer{\'e}di's celebrated Regularity Lemma for graphs interacts with core model-theoretic ideas and techniques. Our work relies on a coincidence of ideas from model theory and graph theory: arbitrarily large half-graphs coincide with model-theoretic instability, so in their absence, structure theorems and technology from stability theory apply. In one direction, we address a problem from the classical Szemer{\'e}di theory. It was known that the "irregular pairs" in the statement of Szemer{\'e}di's Regularity Lemma cannot be eliminated, due to the counterexample of halfgraphs (i.e., the order property, corresponding to model-theoretic instability). We show that half-graphs are the only essential difficulty, by giving a much stronger version of Szemer{\'e}di's Regularity Lemma for models of stable theories of graphs (i.e. graphs with the non-k*-order property), in which there are no irregular pairs, the bounds are significantly improved, and each component satisfies an indivisibility condition. In the other direction, we take a more model-theoretic approach, and give several new Szemer{\'e}di-type partition theorems for models of stable theories of graphs. The first theorem gives a partition of any such graph into indiscernible components, meaning here that each component is either a complete or an empty graph, whose interaction is strongly uniform. This relies on a finitary version of the classic model-theoretic fact that stable theories admit large sets of indiscernibles, by showing that in models of stable theories of graphs one can extract much larger indiscernible sets than expected by Ramsey's theorem. The second and third theorems allow for a much smaller number of components at the cost of weakening the "indivisibility" condition on the components. We also discuss some extensions to graphs without the independence property. All graphs are finite and all partitions are equitable, i.e. the sizes of the components differ by at most 1. In the last three theorems, the number of components depends on the size of the graph; in the first theorem quoted, this number is a function of ε only as in the usual Szemer{\'e}di Regularity Lemma.}, + author = {M. Malliaris And S. Shelah}, + date-added = {2021-11-04 20:42:00 +0000}, + date-modified = {2021-11-04 20:42:00 +0000}, + issn = {00029947}, + journal = {Transactions of the American Mathematical Society}, + number = {3}, + pages = {1551-1585}, + publisher = {American Mathematical Society}, + title = {Regularity Lemmas for Stable Graphs}, + url = {http://www.jstor.org/stable/23813167}, + volume = {366}, + year = {2014}, + bdsk-url-1 = {http://www.jstor.org/stable/23813167}} diff --git a/thys/Szemeredi_Regularity/document/root.tex b/thys/Szemeredi_Regularity/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Szemeredi_Regularity/document/root.tex @@ -0,0 +1,40 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} +\usepackage[english]{babel} % for guillemots + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Szemerédi's Regularity Lemma} +\author{Chelsea Edmonds, Angeliki Koutsoukou-Argyraki and Lawrence C. Paulson\\ +Computer Laboratory, University of Cambridge CB3 0FD\\ +\texttt{\{cle47,ak2110,lp15\}@cam.ac.uk}} + +\maketitle + +\begin{abstract} +Szemerédi's regularity lemma \cite{szemeredi-regular} is a key result in the study of large graphs. It asserts the existence an upper bound on the number of parts the vertices of a graph need to be partitioned into such that the edges between the parts are random in a certain sense. This bound depends only on the desired precision and not on the graph itself, in the spirit of Ramsey's theorem. The formalisation follows online course notes by Tim Gowers\footnote{\url{https://www.dpmms.cam.ac.uk/~par31/notes/tic.pdf}} +and Yufei Zhao\footnote{\url{https://yufeizhao.com/gtac/gtac.pdf} and \url{https://yufeizhao.com/gtac/gtac17.pdf} are successive drafts of a textbook in preparation.}. +Similar material is found in many textbooks \cite{diestel-graph}. +\end{abstract} + +\tableofcontents + +\subsection*{Acknowledgements} +The authors were supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council. + +\bibliographystyle{abbrv} +\bibliography{root} +\newpage + +% include generated text of all theories +\input{session} + +\end{document} diff --git a/thys/Word_Lib/Aligned.thy b/thys/Word_Lib/Aligned.thy --- a/thys/Word_Lib/Aligned.thy +++ b/thys/Word_Lib/Aligned.thy @@ -1,1325 +1,1325 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports "HOL-Library.Word" More_Word Bit_Shifts_Infix_Syntax begin context includes bit_operations_syntax begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ is \\k n. 2 ^ n dvd take_bit LENGTH('a) k\ by simp lemma is_aligned_iff_udvd: \is_aligned w n \ 2 ^ n udvd w\ by transfer (simp flip: take_bit_eq_0_iff add: min_def) lemma is_aligned_iff_take_bit_eq_0: \is_aligned w n \ take_bit n w = 0\ by (simp add: is_aligned_iff_udvd take_bit_eq_0_iff exp_dvd_iff_exp_udvd) lemma is_aligned_iff_dvd_int: \is_aligned ptr n \ 2 ^ n dvd uint ptr\ by transfer simp lemma is_aligned_iff_dvd_nat: \is_aligned ptr n \ 2 ^ n dvd unat ptr\ proof - have \unat ptr = nat \uint ptr\\ by transfer simp then have \2 ^ n dvd unat ptr \ 2 ^ n dvd uint ptr\ by (simp only: dvd_nat_abs_iff) simp then show ?thesis by (simp add: is_aligned_iff_dvd_int) qed lemma is_aligned_0 [simp]: \is_aligned 0 n\ by transfer simp lemma is_aligned_at_0 [simp]: \is_aligned w 0\ by transfer simp lemma is_aligned_beyond_length: \is_aligned w n \ w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that apply (simp add: is_aligned_iff_udvd) apply transfer apply auto done lemma is_alignedI [intro?]: \is_aligned x n\ if \x = 2 ^ n * k\ for x :: \'a::len word\ proof (unfold is_aligned_iff_udvd) from that show \2 ^ n udvd x\ using dvd_triv_left exp_dvd_iff_exp_udvd by blast qed lemma is_alignedE: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ proof (cases \n < LENGTH('a)\) case False with assms have \w = 0\ by (simp add: is_aligned_beyond_length) with that [of 0] show thesis by simp next case True moreover define m where \m = LENGTH('a) - n\ ultimately have l: \LENGTH('a) = n + m\ and \m \ 0\ by simp_all from \n < LENGTH('a)\ have *: \unat (2 ^ n :: 'a word) = 2 ^ n\ by transfer simp from assms have \2 ^ n udvd w\ by (simp add: is_aligned_iff_udvd) then obtain v :: \'a word\ where \unat w = unat (2 ^ n :: 'a word) * unat v\ .. moreover define q where \q = unat v\ ultimately have unat_w: \unat w = 2 ^ n * q\ by (simp add: *) then have \word_of_nat (unat w) = (word_of_nat (2 ^ n * q) :: 'a word)\ by simp then have w: \w = 2 ^ n * word_of_nat q\ by simp moreover have \q < 2 ^ (LENGTH('a) - n)\ proof (rule ccontr) assume \\ q < 2 ^ (LENGTH('a) - n)\ then have \2 ^ (LENGTH('a) - n) \ q\ by simp then have \2 ^ LENGTH('a) \ 2 ^ n * q\ by (simp add: l power_add) with unat_w [symmetric] show False - by (metis le_antisym nat_less_le unsigned_less) + by (metis le_antisym nat_less_le unsigned_less) qed ultimately show thesis using that by blast qed lemma is_alignedE' [elim?]: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ proof - from assms obtain q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ by (rule is_alignedE) then have \w = push_bit n (word_of_nat q)\ by (simp add: push_bit_eq_mult) with that show thesis using \q < 2 ^ (LENGTH('a) - n)\ . qed lemma is_aligned_mask: \is_aligned w n \ w AND mask n = 0\ by (simp add: is_aligned_iff_take_bit_eq_0 take_bit_eq_mask) lemma is_aligned_imp_not_bit: \\ bit w m\ if \is_aligned w n\ and \m < n\ for w :: \'a::len word\ proof - from \is_aligned w n\ obtain q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ .. moreover have \\ bit (push_bit n (word_of_nat q :: 'a word)) m\ using \m < n\ by (simp add: bit_simps) ultimately show ?thesis by simp qed lemma is_aligned_weaken: "\ is_aligned w x; x \ y \ \ is_aligned w y" unfolding is_aligned_iff_dvd_nat by (erule dvd_trans [rotated]) (simp add: le_imp_power_dvd) lemma is_alignedE_pre: fixes w::"'a::len word" assumes aligned: "is_aligned w n" shows rl: "\q. w = 2 ^ n * (of_nat q) \ q < 2 ^ (LENGTH('a) - n)" using aligned is_alignedE by blast lemma aligned_add_aligned: fixes x::"'a::len word" assumes aligned1: "is_aligned x n" and aligned2: "is_aligned y m" and lt: "m \ n" shows "is_aligned (x + y) m" proof cases assume nlt: "n < LENGTH('a)" show ?thesis unfolding is_aligned_iff_dvd_nat dvd_def proof - from aligned2 obtain q2 where yv: "y = 2 ^ m * of_nat q2" and q2v: "q2 < 2 ^ (LENGTH('a) - m)" by (auto elim: is_alignedE) from lt obtain k where kv: "m + k = n" by (auto simp: le_iff_add) with aligned1 obtain q1 where xv: "x = 2 ^ (m + k) * of_nat q1" and q1v: "q1 < 2 ^ (LENGTH('a) - (m + k))" by (auto elim: is_alignedE) have l1: "2 ^ (m + k) * q1 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q1v]) (subst kv, rule order_less_imp_le [OF nlt]) have l2: "2 ^ m * q2 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q2v], rule order_less_imp_le [OF order_le_less_trans]) fact+ have "x = of_nat (2 ^ (m + k) * q1)" using xv by simp moreover have "y = of_nat (2 ^ m * q2)" using yv by simp ultimately have upls: "unat x + unat y = 2 ^ m * (2 ^ k * q1 + q2)" proof - have f1: "unat x = 2 ^ (m + k) * q1" using q1v unat_mult_power_lem xv by blast have "unat y = 2 ^ m * q2" using q2v unat_mult_power_lem yv by blast then show ?thesis using f1 by (simp add: power_add semiring_normalization_rules(34)) qed (* (2 ^ k * q1 + q2) *) show "\d. unat (x + y) = 2 ^ m * d" proof (cases "unat x + unat y < 2 ^ LENGTH('a)") case True have "unat (x + y) = unat x + unat y" by (subst unat_plus_if', rule if_P) fact also have "\ = 2 ^ m * (2 ^ k * q1 + q2)" by (rule upls) finally show ?thesis .. next case False then have "unat (x + y) = (unat x + unat y) mod 2 ^ LENGTH('a)" by (subst unat_word_ariths(1)) simp also have "\ = (2 ^ m * (2 ^ k * q1 + q2)) mod 2 ^ LENGTH('a)" by (subst upls, rule refl) also have "\ = 2 ^ m * ((2 ^ k * q1 + q2) mod 2 ^ (LENGTH('a) - m))" proof - have "m \ len_of (TYPE('a))" by (meson le_trans less_imp_le_nat lt nlt) then show ?thesis by (metis mult_mod_right ordered_cancel_comm_monoid_diff_class.add_diff_inverse power_add) qed finally show ?thesis .. qed qed next assume "\ n < LENGTH('a)" with assms show ?thesis by (simp add: is_aligned_mask not_less take_bit_eq_mod power_overflow word_arith_nat_defs(7) flip: take_bit_eq_mask) qed corollary aligned_sub_aligned: "\is_aligned (x::'a::len word) n; is_aligned y m; m \ n\ \ is_aligned (x - y) m" apply (simp del: add_uminus_conv_diff add:diff_conv_add_uminus) apply (erule aligned_add_aligned, simp_all) apply (erule is_alignedE) apply (rule_tac k="- of_nat q" in is_alignedI) apply simp done lemma is_aligned_shift: fixes k::"'a::len word" shows "is_aligned (k << m) m" proof cases assume mv: "m < LENGTH('a)" from mv obtain q where mq: "m + q = LENGTH('a)" and "0 < q" by (auto dest: less_imp_add_positive) have "(2::nat) ^ m dvd unat (push_bit m k)" proof have kv: "(unat k div 2 ^ q) * 2 ^ q + unat k mod 2 ^ q = unat k" by (rule div_mult_mod_eq) have "unat (push_bit m k) = unat (2 ^ m * k)" by (simp add: push_bit_eq_mult ac_simps) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (simp add: unat_word_ariths(2)) also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp finally show "unat (push_bit m k) = 2 ^ m * (unat k mod 2 ^ q)" . qed then show ?thesis by (unfold is_aligned_iff_dvd_nat shiftl_def) next assume "\ m < LENGTH('a)" then show ?thesis by (simp add: not_less power_overflow is_aligned_mask word_size shiftl_def) qed lemma aligned_mod_eq_0: fixes p::"'a::len word" assumes al: "is_aligned p sz" shows "p mod 2 ^ sz = 0" proof cases assume szv: "sz < LENGTH('a)" with al show ?thesis unfolding is_aligned_iff_dvd_nat by (simp add: and_mask_dvd_nat p2_gt_0 word_mod_2p_is_mask) next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: is_aligned_mask flip: take_bit_eq_mask take_bit_eq_mod) qed lemma is_aligned_triv: "is_aligned (2 ^ n ::'a::len word) n" by (rule is_alignedI [where k = 1], simp) lemma is_aligned_mult_triv1: "is_aligned (2 ^ n * x ::'a::len word) n" by (rule is_alignedI [OF refl]) lemma is_aligned_mult_triv2: "is_aligned (x * 2 ^ n ::'a::len word) n" by (subst mult.commute, simp add: is_aligned_mult_triv1) lemma word_power_less_0_is_0: fixes x :: "'a::len word" shows "x < a ^ 0 \ x = 0" by simp lemma is_aligned_no_wrap: fixes off :: "'a::len word" fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "unat ptr + unat off < 2 ^ LENGTH('a)" proof - have szv: "sz < LENGTH('a)" using off p2_gt_0 word_neq_0_conv by fastforce from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by simp next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q ::'a::len word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply (simp_all) done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + unat off < 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_less [OF _ qv]) apply (rule order_less_le_trans [OF unat_mono [OF off] order_eq_refl]) apply simp_all done qed qed qed lemma is_aligned_no_wrap': fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "ptr \ ptr + off" by (subst no_plus_overflow_unat_size, subst word_size, rule is_aligned_no_wrap) fact+ lemma is_aligned_no_overflow': fixes p :: "'a::len word" assumes al: "is_aligned p n" shows "p \ p + (2 ^ n - 1)" proof cases assume "n n ptr \ ptr + 2^sz - 1" by (drule is_aligned_no_overflow') (simp add: field_simps) lemma replicate_not_True: "\n. xs = replicate n False \ True \ set xs" by (induct xs) auto lemma map_zip_replicate_False_xor: "n = length xs \ map (\(x, y). x = (\ y)) (zip xs (replicate n False)) = xs" by (induct xs arbitrary: n, auto) lemma drop_minus_lem: "\ n \ length xs; 0 < n; n' = length xs \ \ drop (n' - n) xs = rev xs ! (n - 1) # drop (Suc (n' - n)) xs" proof (induct xs arbitrary: n n') case Nil then show ?case by simp next case (Cons y ys) from Cons.prems show ?case apply simp apply (cases "n = Suc (length ys)") apply (simp add: nth_append) apply (simp add: Suc_diff_le Cons.hyps nth_append) apply clarsimp apply arith done qed lemma drop_minus: "\ n < length xs; n' = length xs \ \ drop (n' - Suc n) xs = rev xs ! n # drop (n' - n) xs" apply (subst drop_minus_lem) apply simp apply simp apply simp apply simp apply (cases "length xs", simp) apply (simp add: Suc_diff_le) done lemma aligned_add_xor: \(x + 2 ^ n) XOR 2 ^ n = x\ if al: \is_aligned (x::'a::len word) n'\ and le: \n < n'\ proof - have \\ bit x n\ using that by (rule is_aligned_imp_not_bit) then have \x + 2 ^ n = x OR 2 ^ n\ by (subst disjunctive_add) (auto simp add: bit_simps disjunctive_add) moreover have \(x OR 2 ^ n) XOR 2 ^ n = x\ by (rule bit_word_eqI) (auto simp add: bit_simps \\ bit x n\) ultimately show ?thesis by simp qed lemma is_aligned_add_mult_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n * z) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x*z"]) done lemma is_aligned_add_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x"]) done lemma is_aligned_no_wrap''': fixes ptr :: "'a::len word" shows"\ is_aligned ptr sz; sz < LENGTH('a); off < 2 ^ sz \ \ unat ptr + off < 2 ^ LENGTH('a)" apply (drule is_aligned_no_wrap[where off="of_nat off"]) apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod unsigned_of_nat) apply (subst(asm) unat_of_nat_len) apply (erule order_less_trans) apply (erule power_strict_increasing) apply simp apply assumption done lemma is_aligned_get_word_bits: fixes p :: "'a::len word" shows "\ is_aligned p n; \ is_aligned p n; n < LENGTH('a) \ \ P; \ p = 0; n \ LENGTH('a) \ \ P \ \ P" apply (cases "n < LENGTH('a)") apply simp apply simp apply (erule meta_mp) apply (simp add: is_aligned_mask power_add power_overflow not_less flip: take_bit_eq_mask) apply (metis take_bit_length_eq take_bit_of_0 take_bit_tightened) done lemma aligned_small_is_0: "\ is_aligned x n; x < 2 ^ n \ \ x = 0" by (simp add: is_aligned_mask less_mask_eq) corollary is_aligned_less_sz: "\is_aligned a sz; a \ 0\ \ \ a < 2 ^ sz" by (rule notI, drule(1) aligned_small_is_0, erule(1) notE) lemma aligned_at_least_t2n_diff: "\is_aligned x n; is_aligned y n; x < y\ \ x \ y - 2 ^ n" apply (erule is_aligned_get_word_bits[where p=y]) apply (rule ccontr) apply (clarsimp simp: linorder_not_le) apply (subgoal_tac "y - x = 0") apply clarsimp apply (rule aligned_small_is_0) apply (erule(1) aligned_sub_aligned) apply simp apply unat_arith apply simp done lemma is_aligned_no_overflow'': "\is_aligned x n; x + 2 ^ n \ 0\ \ x \ x + 2 ^ n" apply (frule is_aligned_no_overflow') apply (erule order_trans) apply (simp add: field_simps) apply (erule word_sub_1_le) done lemma is_aligned_bitI: \is_aligned p m\ if \\n. n < m \ \ bit p n\ apply (simp add: is_aligned_mask) apply (rule bit_word_eqI) using that apply (auto simp add: bit_simps) done lemma is_aligned_nth: "is_aligned p m = (\n < m. \ bit p n)" apply (auto intro: is_aligned_bitI simp add: is_aligned_mask bit_eq_iff) apply (auto simp: bit_simps) using bit_imp_le_length not_less apply blast done lemma range_inter: "({a..b} \ {c..d} = {}) = (\x. \(a \ x \ x \ b \ c \ x \ x \ d))" by auto lemma aligned_inter_non_empty: "\ {p..p + (2 ^ n - 1)} \ {p..p + 2 ^ m - 1} = {}; is_aligned p n; is_aligned p m\ \ False" apply (clarsimp simp only: range_inter) apply (erule_tac x=p in allE) apply simp apply (erule impE) apply (erule is_aligned_no_overflow') apply (erule notE) apply (erule is_aligned_no_overflow) done lemma not_aligned_mod_nz: assumes al: "\ is_aligned a n" shows "a mod 2 ^ n \ 0" apply (rule ccontr) using al apply (rule notE) apply simp apply (rule is_alignedI [of _ _ \a div 2 ^ n\]) apply (metis add.right_neutral mult.commute word_mod_div_equality) done lemma nat_add_offset_le: fixes x :: nat assumes yv: "y \ 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y \ 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" by (auto simp: le_iff_add) have "x * 2 ^ n + y \ x * 2 ^ n + 2 ^ n" using yv xv by simp also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y \ 2 ^ (m + n)" . qed lemma is_aligned_no_wrap_le: fixes ptr::"'a::len word" assumes al: "is_aligned ptr sz" and szv: "sz < LENGTH('a)" and off: "off \ 2 ^ sz" shows "unat ptr + off \ 2 ^ LENGTH('a)" proof - from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by (auto simp add: le_Suc_eq Suc_le_eq) (simp add: le_less) next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q :: 'a word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply simp_all done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + off \ 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_le [OF off qv]) apply simp done qed qed qed lemma is_aligned_neg_mask: "m \ n \ is_aligned (x AND NOT (mask n)) m" by (rule is_aligned_bitI) (simp add: bit_simps) lemma unat_minus: "unat (- (x :: 'a :: len word)) = (if x = 0 then 0 else 2 ^ size x - unat x)" using unat_sub_if_size[where x="2 ^ size x" and y=x] by (simp add: unat_eq_0 word_size) lemma is_aligned_minus: \is_aligned (- p) n\ if \is_aligned p n\ for p :: \'a::len word\ using that apply (cases \n < LENGTH('a)\) apply (simp_all add: not_less is_aligned_beyond_length) apply transfer apply (simp flip: take_bit_eq_0_iff) apply (subst take_bit_minus [symmetric]) apply simp done lemma add_mask_lower_bits: "\is_aligned (x :: 'a :: len word) n; \n' \ n. n' < LENGTH('a) \ \ bit p n'\ \ x + p AND NOT (mask n) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth) apply (erule_tac x=na in allE)+ apply (simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_less word_size) apply (metis is_aligned_nth not_le) done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (w << m) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (w >> m) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl_self: "is_aligned (p << n) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p AND NOT (mask n) = p" apply (rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) done lemma is_aligned_shiftr_shiftl: "is_aligned w n \ w >> n << n = w" apply (rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth intro: ccontr) apply (subst add_diff_inverse_nat) apply (auto intro: ccontr) done lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ ((x >> n) AND mask v) << n = x AND mask (v + n)" apply (rule word_eqI) apply (simp add: word_size bit_simps) apply (subgoal_tac "\m. bit x m \ m \ n") apply auto[1] apply (clarsimp simp: is_aligned_mask) apply (drule_tac x=m in word_eqD) apply (frule test_bit_size) apply (simp add: word_size bit_simps) done lemma mask_zero: "is_aligned x a \ x AND mask a = 0" by (metis is_aligned_mask) lemma is_aligned_neg_mask_eq_concrete: "\ is_aligned p n; msk AND NOT (mask n) = NOT (mask n) \ \ p AND msk = p" by (metis word_bw_assocs(1) word_bw_comms(1) is_aligned_neg_mask_eq) lemma is_aligned_and_not_zero: "\ is_aligned n k; n \ 0 \ \ 2 ^ k \ n" using is_aligned_less_sz leI by blast lemma is_aligned_and_2_to_k: "(n AND 2 ^ k - 1) = 0 \ is_aligned (n :: 'a :: len word) k" by (simp add: is_aligned_mask mask_eq_decr_exp) lemma is_aligned_power2: "b \ a \ is_aligned (2 ^ a) b" by (metis is_aligned_triv is_aligned_weaken) lemma aligned_sub_aligned': "\ is_aligned (a :: 'a :: len word) n; is_aligned b n; n < LENGTH('a) \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma is_aligned_neg_mask_weaken: "\ is_aligned p n; m \ n \ \ p AND NOT (mask m) = p" using is_aligned_neg_mask_eq is_aligned_weaken by blast lemma is_aligned_neg_mask2 [simp]: "is_aligned (a AND NOT (mask n)) n" by (rule is_aligned_bitI) (simp add: bit_simps) lemma is_aligned_0': "is_aligned 0 n" by (fact is_aligned_0) lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply (simp flip: take_bit_eq_mod) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) apply (metis assms(2) bit_or_iff is_aligned_mask is_aligned_nth leD less_mask_eq word_and_le1 word_bw_lcs(1) word_neq_0_conv word_plus_and_or_coroll) apply (meson assms(2) leI less_2p_is_upper_bits_unset) apply (metis assms(2) bit_disjunctive_add_iff bit_imp_le_length bit_push_bit_iff is_alignedE' less_2p_is_upper_bits_unset) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_eq_decr_exp word_mod_by_0) qed lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply transfer apply simp apply (auto simp add: le_less power_2_ge_iff szv) apply (metis le_less_trans mask_eq_decr_exp mask_lt_2pn order_less_imp_le szv) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_eq_decr_exp power_overflow) then show ?thesis .. qed lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d AND mask n = d) \ (p + d AND (NOT (mask n)) = p)" apply (subst (asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll) apply (simp_all flip: take_bit_eq_mask) apply (metis take_bit_eq_mask word_bw_lcs(1) word_log_esimps(1)) apply (metis add.commute add_left_imp_eq take_bit_eq_mask word_plus_and_or_coroll2) done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) AND mask n = q AND mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q AND NOT (mask n)) = (p + q) AND NOT (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p OR d" apply (subst disjunctive_add, simp_all) apply (clarsimp simp: is_aligned_nth less_2p_is_upper_bits_unset) subgoal for m apply (cases \m < n\) apply (auto simp add: not_less dest: bit_imp_possible_bit) done done lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" by (metis NOT_mask add_diff_cancel_right' diff_0 is_aligned_neg_mask_eq not_le word_and_le1) lemma neg_mask_mono_le: "x \ y \ x AND NOT(mask n) \ y AND NOT(mask n)" for x :: "'a :: len word" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False then show "y AND NOT(mask n) < x AND NOT(mask n) \ False" by (simp add: mask_eq_decr_exp linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y AND NOT(mask n) < x AND NOT(mask n)" have word_bits: "n < LENGTH('a)" by fact have "y \ (y AND NOT(mask n)) + (y AND mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y AND NOT(mask n)) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b; simp add: is_aligned_neg_mask) done also have "\ \ x AND NOT(mask n)" using b apply (subst add.commute) apply (rule le_plus) apply (rule aligned_at_least_t2n_diff; simp add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated]; simp add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma and_neg_mask_eq_iff_not_mask_le: "w AND NOT(mask n) = NOT(mask n) \ NOT(mask n) \ w" for w :: \'a::len word\ by (metis eq_iff neg_mask_mono_le word_and_le1 word_and_le2 word_bw_same(1)) lemma neg_mask_le_high_bits: \NOT (mask n) \ w \ (\i \ {n ..< size w}. bit w i)\ (is \?P \ ?Q\) for w :: \'a::len word\ proof assume ?Q then have \w AND NOT (mask n) = NOT (mask n)\ by (auto simp add: bit_simps word_size intro: bit_word_eqI) then show ?P by (simp add: and_neg_mask_eq_iff_not_mask_le) next assume ?P then have *: \w AND NOT (mask n) = NOT (mask n)\ by (simp add: and_neg_mask_eq_iff_not_mask_le) show \?Q\ proof (rule ccontr) assume \\ (\i\{n.. then obtain m where m: \\ bit w m\ \n \ m\ \m < LENGTH('a)\ by (auto simp add: word_size) from * have \bit (w AND NOT (mask n)) m \ bit (NOT (mask n :: 'a word)) m\ by auto with m show False by (auto simp add: bit_simps) qed qed lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_eq_decr_exp) (erule is_aligned_no_overflow') lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: dvd_def is_aligned_iff_dvd_nat) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" - by (meson is_aligned_get_word_bits) + by (meson is_aligned_get_word_bits) thus ?thesis by simp qed lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less not_le) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (subst (asm) unat_add_lem' [symmetric]) apply (simp add: is_aligned_mask_offset_unat) apply (metis gap_between_aligned linorder_not_less mask_eq_decr_exp unat_arith_simps(2)) done lemma aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: of_nat_diff) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" using take_bit_nat_less_self_iff [of \LENGTH('a)\ \mq - 2 ^ sq * nq\] apply (auto simp add: word_less_nat_alt not_le not_less unsigned_of_nat) apply (metis take_bit_nat_eq_self_iff) done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt take_bit_eq_mod unsigned_of_nat) apply (subst (asm) mod_less) apply auto apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp done ultimately show ?thesis by auto qed lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ (x + y) >> n = y >> n" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) apply (metis less_2p_is_upper_bits_unset not_le) apply (metis le_add1 less_2p_is_upper_bits_unset test_bit_bin) done lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ (y + x) >> n = y >> n" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) apply (metis less_2p_is_upper_bits_unset not_le) apply (metis bit_imp_le_length le_add1 less_2p_is_upper_bits_unset) done lemma and_neg_mask_plus_mask_mono: "(p AND NOT (mask n)) + mask n \ p" for p :: \'a::len word\ apply (rule word_le_minus_cancel[where x = "p AND NOT (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_eq_decr_exp word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr AND NOT (mask n)) + (2 ^ n - 1)" for ptr :: \'a::len word\ by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p AND mask n = d) \ (p AND (NOT (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k AND mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a AND mask n = (ptr AND mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a AND mask n) AND NOT (mask m) = (ptr + a AND NOT (mask m) ) AND mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (meson nat_mult_power_less_eq zero_less_numeral) done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) apply transfer apply (auto simp add: is_aligned_iff_udvd) apply (metis (no_types, lifting) le_less_trans less_not_refl2 less_or_eq_imp_le of_nat_eq_0_iff take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat unat_lt2p unat_power_lower) done lemma le_or_mask: "w \ w' \ w OR mask x \ w' OR mask x" for w w' :: \'a::len word\ by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) end end diff --git a/thys/Word_Lib/Bit_Comprehension.thy b/thys/Word_Lib/Bit_Comprehension.thy --- a/thys/Word_Lib/Bit_Comprehension.thy +++ b/thys/Word_Lib/Bit_Comprehension.thy @@ -1,251 +1,251 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Comprehension syntax for bit expressions\ theory Bit_Comprehension imports "HOL-Library.Word" begin class bit_comprehension = ring_bit_operations + fixes set_bits :: \(nat \ bool) \ 'a\ (binder \BITS \ 10) assumes set_bits_bit_eq: \set_bits (bit a) = a\ begin lemma set_bits_False_eq [simp]: \(BITS _. False) = 0\ using set_bits_bit_eq [of 0] by (simp add: bot_fun_def) end instantiation int :: bit_comprehension begin definition \set_bits f = ( if \n. \m\n. f m = f n then let n = LEAST n. \m\n. f m = f n in signed_take_bit n (horner_sum of_bool 2 (map f [0.. instance proof fix k :: int from int_bit_bound [of k] obtain n where *: \\m. n \ m \ bit k m \ bit k n\ and **: \n > 0 \ bit k (n - 1) \ bit k n\ by blast then have ***: \\n. \n'\n. bit k n' \ bit k n\ by meson have l: \(LEAST q. \m\q. bit k m \ bit k q) = n\ apply (rule Least_equality) using * apply blast apply (metis "**" One_nat_def Suc_pred le_cases le0 neq0_conv not_less_eq_eq) done show \set_bits (bit k) = k\ apply (simp only: *** set_bits_int_def horner_sum_bit_eq_take_bit l) apply simp apply (rule bit_eqI) apply (simp add: bit_signed_take_bit_iff min_def) apply (auto simp add: not_le bit_take_bit_iff dest: *) done qed end lemma int_set_bits_K_False [simp]: "(BITS _. False) = (0 :: int)" by (simp add: set_bits_int_def) lemma int_set_bits_K_True [simp]: "(BITS _. True) = (-1 :: int)" by (simp add: set_bits_int_def) instantiation word :: (len) bit_comprehension begin definition word_set_bits_def: \(BITS n. P n) = (horner_sum of_bool 2 (map P [0.. instance by standard (simp add: word_set_bits_def horner_sum_bit_eq_take_bit) end lemma bit_set_bits_word_iff [bit_simps]: \bit (set_bits P :: 'a::len word) n \ n < LENGTH('a) \ P n\ by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) lemma set_bits_K_False [simp]: \set_bits (\_. False) = (0 :: 'a :: len word)\ by (rule bit_word_eqI) (simp add: bit_set_bits_word_iff) lemma set_bits_int_unfold': \set_bits f = (if \n. \n'\n. \ f n' then let n = LEAST n. \n'\n. \ f n' in horner_sum of_bool 2 (map f [0..n. \n'\n. f n' then let n = LEAST n. \n'\n. f n' in signed_take_bit n (horner_sum of_bool 2 (map f [0.. proof (cases \\n. \m\n. f m \ f n\) case True then obtain q where q: \\m\q. f m \ f q\ by blast define n where \n = (LEAST n. \m\n. f m \ f n)\ have \\m\n. f m \ f n\ unfolding n_def using q by (rule LeastI [of _ q]) then have n: \\m. n \ m \ f m \ f n\ by blast from n_def have n_eq: \(LEAST q. \m\q. f m \ f n) = n\ - by (smt Least_equality Least_le \\m\n. f m = f n\ dual_order.refl le_refl n order_refl) + by (smt (verit, best) Least_le \\m\n. f m = f n\ dual_order.antisym wellorder_Least_lemma(1)) show ?thesis proof (cases \f n\) case False with n have *: \\n. \n'\n. \ f n'\ by blast have **: \(LEAST n. \n'\n. \ f n') = n\ using False n_eq by simp from * False show ?thesis apply (simp add: set_bits_int_def n_def [symmetric] ** del: upt.upt_Suc) apply (auto simp add: take_bit_horner_sum_bit_eq bit_horner_sum_bit_iff take_map signed_take_bit_def set_bits_int_def horner_sum_bit_eq_take_bit simp del: upt.upt_Suc) done next case True with n have *: \\n. \n'\n. f n'\ by blast have ***: \\ (\n. \n'\n. \ f n')\ apply (rule ccontr) using * nat_le_linear by auto have **: \(LEAST n. \n'\n. f n') = n\ using True n_eq by simp from * *** True show ?thesis apply (simp add: set_bits_int_def n_def [symmetric] ** del: upt.upt_Suc) apply (auto simp add: take_bit_horner_sum_bit_eq bit_horner_sum_bit_iff take_map signed_take_bit_def set_bits_int_def horner_sum_bit_eq_take_bit nth_append simp del: upt.upt_Suc) done qed next case False then show ?thesis by (auto simp add: set_bits_int_def) qed inductive wf_set_bits_int :: "(nat \ bool) \ bool" for f :: "nat \ bool" where zeros: "\n' \ n. \ f n' \ wf_set_bits_int f" | ones: "\n' \ n. f n' \ wf_set_bits_int f" lemma wf_set_bits_int_simps: "wf_set_bits_int f \ (\n. (\n'\n. \ f n') \ (\n'\n. f n'))" by(auto simp add: wf_set_bits_int.simps) lemma wf_set_bits_int_const [simp]: "wf_set_bits_int (\_. b)" by(cases b)(auto intro: wf_set_bits_int.intros) lemma wf_set_bits_int_fun_upd [simp]: "wf_set_bits_int (f(n := b)) \ wf_set_bits_int f" (is "?lhs \ ?rhs") proof assume ?lhs then obtain n' where "(\n''\n'. \ (f(n := b)) n'') \ (\n''\n'. (f(n := b)) n'')" by(auto simp add: wf_set_bits_int_simps) hence "(\n''\max (Suc n) n'. \ f n'') \ (\n''\max (Suc n) n'. f n'')" by auto thus ?rhs by(auto simp only: wf_set_bits_int_simps) next assume ?rhs then obtain n' where "(\n''\n'. \ f n'') \ (\n''\n'. f n'')" (is "?wf f n'") by(auto simp add: wf_set_bits_int_simps) hence "?wf (f(n := b)) (max (Suc n) n')" by auto thus ?lhs by(auto simp only: wf_set_bits_int_simps) qed lemma wf_set_bits_int_Suc [simp]: "wf_set_bits_int (\n. f (Suc n)) \ wf_set_bits_int f" (is "?lhs \ ?rhs") by(auto simp add: wf_set_bits_int_simps intro: le_SucI dest: Suc_le_D) context fixes f assumes wff: "wf_set_bits_int f" begin lemma int_set_bits_unfold_BIT: "set_bits f = of_bool (f 0) + (2 :: int) * set_bits (f \ Suc)" using wff proof cases case (zeros n) show ?thesis proof(cases "\n. \ f n") case True hence "f = (\_. False)" by auto thus ?thesis using True by(simp add: o_def) next case False then obtain n' where "f n'" by blast with zeros have "(LEAST n. \n'\n. \ f n') = Suc (LEAST n. \n'\Suc n. \ f n')" by(auto intro: Least_Suc) also have "(\n. \n'\Suc n. \ f n') = (\n. \n'\n. \ f (Suc n'))" by(auto dest: Suc_le_D) also from zeros have "\n'\n. \ f (Suc n')" by auto ultimately show ?thesis using zeros apply (simp (no_asm_simp) add: set_bits_int_unfold' exI del: upt.upt_Suc flip: map_map split del: if_split) apply (simp only: map_Suc_upt upt_conv_Cons) apply simp done qed next case (ones n) show ?thesis proof(cases "\n. f n") case True hence "f = (\_. True)" by auto thus ?thesis using True by(simp add: o_def) next case False then obtain n' where "\ f n'" by blast with ones have "(LEAST n. \n'\n. f n') = Suc (LEAST n. \n'\Suc n. f n')" by(auto intro: Least_Suc) also have "(\n. \n'\Suc n. f n') = (\n. \n'\n. f (Suc n'))" by(auto dest: Suc_le_D) also from ones have "\n'\n. f (Suc n')" by auto moreover from ones have "(\n. \n'\n. \ f n') = False" by(auto intro!: exI[where x="max n m" for n m] simp add: max_def split: if_split_asm) moreover hence "(\n. \n'\n. \ f (Suc n')) = False" by(auto elim: allE[where x="Suc n" for n] dest: Suc_le_D) ultimately show ?thesis using ones apply (simp (no_asm_simp) add: set_bits_int_unfold' exI split del: if_split) apply (auto simp add: Let_def hd_map map_tl[symmetric] map_map[symmetric] map_Suc_upt upt_conv_Cons signed_take_bit_Suc not_le simp del: map_map) done qed qed lemma bin_last_set_bits [simp]: "odd (set_bits f :: int) = f 0" by (subst int_set_bits_unfold_BIT) simp_all lemma bin_rest_set_bits [simp]: "set_bits f div (2 :: int) = set_bits (f \ Suc)" by (subst int_set_bits_unfold_BIT) simp_all lemma bin_nth_set_bits [simp]: "bit (set_bits f :: int) m \ f m" using wff proof (induction m arbitrary: f) case 0 then show ?case by (simp add: Bit_Comprehension.bin_last_set_bits) next case Suc from Suc.IH [of "f \ Suc"] Suc.prems show ?case by (simp add: Bit_Comprehension.bin_rest_set_bits comp_def bit_Suc) qed end end diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy --- a/thys/Word_Lib/Bits_Int.thy +++ b/thys/Word_Lib/Bits_Int.thy @@ -1,1573 +1,1577 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Bitwise Operations on integers\ theory Bits_Int imports "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" "Word_Lib.Bit_Comprehension" begin subsection \Implicit bit representation of \<^typ>\int\\ lemma bin_last_def: "(odd :: int \ bool) w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) lemma bin_last_numeral_simps [simp]: "\ odd (0 :: int)" "odd (1 :: int)" "odd (- 1 :: int)" "odd (Numeral1 :: int)" "\ odd (numeral (Num.Bit0 w) :: int)" "odd (numeral (Num.Bit1 w) :: int)" "\ odd (- numeral (Num.Bit0 w) :: int)" "odd (- numeral (Num.Bit1 w) :: int)" by simp_all lemma bin_rest_numeral_simps [simp]: "(\k::int. k div 2) 0 = 0" "(\k::int. k div 2) 1 = 0" "(\k::int. k div 2) (- 1) = - 1" "(\k::int. k div 2) Numeral1 = 0" "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd y\ \ x = y" by (auto elim: oddE) lemma [simp]: shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) x > 0 \ x > 1" by auto subsection \Bit projection\ lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" using that by (rule bit_eqI) lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" by (metis bit_eq_iff) lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) lemma bin_nth_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" by simp lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps lemma nth_2p_bin: "(bit :: int \ nat \ bool) (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done lemma bin_nth_numeral_unfold: "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) w) = bin_sign w" by (simp add: bin_sign_def) lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) n w = w mod 2 ^ n" by (fact take_bit_eq_mod) lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: \(signed_take_bit :: nat \ int \ int) n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc del: take_bit_minus_one_eq_mask) lemma sbintrunc_0_numeral [simp]: "(signed_take_bit :: nat \ int \ int) 0 1 = -1" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) n bin) = -1) = bit bin n" by (simp add: bin_sign_def) lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) w n" by (fact bit_take_bit_iff) lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by simp lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" by simp lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (n - 1) w" by (cases n) simp_all lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp moreover have \(signed_take_bit :: nat \ int \ int) n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_int_greater_eq) lemma sbintrunc_dec: \(signed_take_bit :: nat \ int \ int) n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_int_less_eq) lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) lemma bintr_lt2p: "(take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: bintrunc_mod2p) lemma bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq mask_eq_exp_minus_1) lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" by (fact signed_take_bit_int_greater_eq_minus_exp) lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" by (fact signed_take_bit_int_less_exp) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (Suc n) bin)" by (auto simp add: take_bit_Suc) lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) lemma bin_cat_eq_push_bit_add_take_bit: \concat_bit n l k = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) y n z)" by (fact concat_bit_assoc) lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: \nat \ int list \ int\ where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ lemma bin_rcat_eq_foldl: \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0 ks\ by (induction ks rule: rev_induct) (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) qed fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) k a) n b" by (metis bin_cat_assoc bin_cat_zero) lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: \drop_bit n ((\k n l. concat_bit n l k) v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: \take_bit n ((\k n l. concat_bit n l k) v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ abbreviation (input) bin_sc :: \nat \ bool \ int \ int\ where \bin_sc n b i \ set_bit i n b\ lemma bin_sc_0 [simp]: "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" by (simp add: set_bit_int_def) lemma bin_sc_Suc [simp]: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" by (simp add: set_bit_int_def set_bit_Suc unset_bit_Suc bin_last_def) lemma bin_nth_sc [bit_simps]: "bit (bin_sc n b w) n \ b" by (simp add: bit_simps) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) done lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ apply (simp_all add: fun_eq_iff bit_eq_iff) apply (simp_all add: bit_simps bin_nth_sc_gen) done lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" apply (rule bit_eqI) apply (cases x) apply (auto simp add: bit_simps bin_sc_eq) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: set_bit_int_def unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: set_bit_int_def set_bit_greater_eq) lemma bintr_bin_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_unset_bit_eq unset_bit_less_eq) lemma bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc_0 bin_sc_Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc_Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma int_set_bit_0 [simp]: fixes x :: int shows "set_bit x 0 b = of_bool b + 2 * (x div 2)" by (fact bin_sc_0) lemma int_set_bit_Suc: fixes x :: int shows "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" by (fact bin_sc_Suc) lemma bin_last_set_bit: "odd (set_bit x n b :: int) = (if n > 0 then odd x else b)" by (cases n) (simp_all add: int_set_bit_Suc) lemma bin_rest_set_bit: "(set_bit x n b :: int) div 2 = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" by (cases n) (simp_all add: int_set_bit_Suc) lemma int_set_bit_numeral: fixes x :: int shows "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" by (fact bin_sc_numeral) lemmas int_set_bit_numerals [simp] = int_set_bit_numeral[where x="numeral w'"] int_set_bit_numeral[where x="- numeral w'"] int_set_bit_numeral[where x="Numeral1"] int_set_bit_numeral[where x="1"] int_set_bit_numeral[where x="0"] int_set_bit_Suc[where x="numeral w'"] int_set_bit_Suc[where x="- numeral w'"] int_set_bit_Suc[where x="Numeral1"] int_set_bit_Suc[where x="1"] int_set_bit_Suc[where x="0"] for w' lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" by (simp add: msb_int_def set_bit_int_def) lemma word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ apply (rule bit_word_eqI) apply (cases x) apply (simp_all add: bit_simps bin_sc_eq) done lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" unfolding word_set_bit_def by (rule word_eqI) (simp add: word_size bin_nth_sc_gen nth_bintr bit_simps) lemma word_set_numeral [simp]: "set_bit (numeral bin::'a::len word) n b = word_of_int (bin_sc n b (numeral bin))" unfolding word_numeral_alt by (rule set_bit_word_of_int) lemma word_set_neg_numeral [simp]: "set_bit (- numeral bin::'a::len word) n b = word_of_int (bin_sc n b (- numeral bin))" unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" unfolding word_0_wi by (rule set_bit_word_of_int) lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" unfolding word_1_wi by (rule set_bit_word_of_int) lemma shiftl_int_def: "push_bit n x = x * 2 ^ n" for x :: int by (fact push_bit_eq_mult) lemma shiftr_int_def: "drop_bit n x = x div 2 ^ n" for x :: int by (fact drop_bit_eq_div) subsubsection \Basic simplification rules\ context includes bit_operations_syntax begin lemmas int_not_def = not_int_def lemma int_not_simps: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (fact and.left_neutral) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (fact or.left_neutral) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (fact xor.left_neutral) subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) y" by (subst xor_int_rec) auto lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) y)" by (subst xor_int_rec) auto lemma bin_nth_ops: "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp subsubsection \Interactions with arithmetic\ lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x OR y)" by simp_all lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) n (NOT x)" by (fact take_bit_not_take_bit) text \Want theorems of the form of \bin_trunc_xor\.\ lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ lemma not_int_cmp_0 [simp]: fixes i :: int shows "0 < NOT i \ i < -1" "0 \ NOT i \ i < 0" "NOT i < 0 \ i \ 0" "NOT i \ 0 \ i \ -1" by(simp_all add: int_not_def) arith+ lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" by (fact bit.conj_disj_distrib) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: "(odd :: int \ bool) i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0: "push_bit 0 x = x" and int_shiftl_Suc: "push_bit (Suc n) x = 2 * push_bit n x" by (auto simp add: shiftl_int_def) lemma int_0_shiftl: "push_bit n 0 = (0 :: int)" by (fact push_bit_of_0) lemma bin_last_shiftl: "odd (push_bit n x) \ n = 0 \ (odd :: int \ bool) x" by simp lemma bin_rest_shiftl: "(\k::int. k div 2) (push_bit n x) = (if n > 0 then push_bit (n - 1) x else (\k::int. k div 2) x)" by (cases n) (simp_all add: push_bit_eq_mult) lemma bin_nth_shiftl: "(bit :: int \ nat \ bool) (push_bit n x) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" by (fact bit_push_bit_iff_int) lemma bin_last_shiftr: "odd (drop_bit n x) \ bit x n" for x :: int by (simp add: bit_iff_odd_drop_bit) lemma bin_rest_shiftr: "(\k::int. k div 2) (drop_bit n x) = drop_bit (Suc n) x" by (simp add: drop_bit_Suc drop_bit_half) lemma bin_nth_shiftr: "(bit :: int \ nat \ bool) (drop_bit n x) m = (bit :: int \ nat \ bool) x (n + m)" by (simp add: bit_simps) lemma bin_nth_conv_AND: fixes x :: int shows "(bit :: int \ nat \ bool) x n \ x AND (push_bit n 1) \ 0" by (fact bit_iff_and_push_bit_not_eq_0) lemma int_shiftl_numeral [simp]: "push_bit (numeral w') (numeral w :: int) = push_bit (pred_numeral w') (numeral (num.Bit0 w))" "push_bit (numeral w') (- numeral w :: int) = push_bit (pred_numeral w') (- numeral (num.Bit0 w))" by(simp_all add: numeral_eq_Suc shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "push_bit (numeral w) (1::int) = push_bit (pred_numeral w) 2" using int_shiftl_numeral [of Num.One w] by (simp only: numeral_eq_Suc push_bit_Suc) simp lemma shiftl_ge_0: fixes i :: int shows "push_bit n i \ 0 \ i \ 0" by (fact push_bit_nonnegative_int_iff) lemma shiftl_lt_0: fixes i :: int shows "push_bit n i < 0 \ i < 0" by (fact push_bit_negative_int_iff) lemma int_shiftl_test_bit: "bit (push_bit i n :: int) m \ m \ i \ bit n (m - i)" by (fact bit_push_bit_iff_int) lemma int_0shiftr: "drop_bit x (0 :: int) = 0" by (fact drop_bit_of_0) lemma int_minus1_shiftr: "drop_bit x (-1 :: int) = -1" by (fact drop_bit_minus_one) lemma int_shiftr_ge_0: fixes i :: int shows "drop_bit n i \ 0 \ i \ 0" by (fact drop_bit_nonnegative_int_iff) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "drop_bit n i < 0 \ i < 0" by (fact drop_bit_negative_int_iff) lemma int_shiftr_numeral [simp]: "drop_bit (numeral w') (1 :: int) = 0" "drop_bit (numeral w') (numeral num.One :: int) = 0" "drop_bit (numeral w') (numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (- numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (- numeral w)" "drop_bit (numeral w') (- numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (- numeral (Num.inc w))" by (simp_all add: numeral_eq_Suc add_One drop_bit_Suc) lemma int_shiftr_numeral_Suc0 [simp]: "drop_bit (Suc 0) (1 :: int) = 0" "drop_bit (Suc 0) (numeral num.One :: int) = 0" "drop_bit (Suc 0) (numeral (num.Bit0 w) :: int) = numeral w" "drop_bit (Suc 0) (numeral (num.Bit1 w) :: int) = numeral w" "drop_bit (Suc 0) (- numeral (num.Bit0 w) :: int) = - numeral w" "drop_bit (Suc 0) (- numeral (num.Bit1 w) :: int) = - numeral (Num.inc w)" by (simp_all add: drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = push_bit n 1" and m: "m < n" and x: "x < y" shows "bit (x - y) m = bit x m" proof - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ by (simp_all add: bin_sign_def push_bit_eq_mult split: if_splits) from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ proof (induction m arbitrary: x n) case 0 then show ?case by simp next case (Suc m) moreover define q where \q = n - 1\ ultimately have n: \n = Suc q\ by simp have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ by simp moreover from Suc.IH [of \x div 2\ q] Suc.prems have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ by (simp add: n) ultimately show ?case by (simp add: bit_Suc n) qed with \y = 2 ^ n\ show ?thesis by simp qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) lemma bin_set_conv_OR: "bin_sc n True i = i OR (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) end subsection \More lemmas on words\ lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" by (simp add: bin_sign_def not_le msb_int_def) lemma msb_bin_sc: "msb (bin_sc n b x) \ msb x" by (simp add: msb_conv_bin_sign) lemma msb_word_def: \msb a \ bin_sign (signed_take_bit (LENGTH('a) - 1) (uint a)) = - 1\ for a :: \'a::len word\ by (simp add: bin_sign_def bit_simps msb_word_iff_bit) lemma word_msb_def: "msb a \ bin_sign (sint a) = - 1" by (simp add: msb_word_def sint_uint) lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], folded uint_word_of_int_eq, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma setBit_no: "Bit_Operations.set_bit n (numeral bin) = word_of_int (bin_sc n True (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma clearBit_no: "unset_bit n (numeral bin) = word_of_int (bin_sc n False (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" by transfer (simp add: drop_bit_take_bit ac_simps) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int ((\k n l. concat_bit n l k) w (size b) (uint b))" by transfer (simp add: take_bit_concat_bit_eq) lemma bintrunc_shiftl: "take_bit n (push_bit i m) = push_bit i (take_bit (n - i) m)" for m :: int by (fact take_bit_push_bit) lemma uint_shiftl: "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))" by (simp add: unsigned_push_bit_eq word_size) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) context includes bit_operations_syntax begin lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) end lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = push_bit n (1 :: int)" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma sint_range': \- (2 ^ (LENGTH('a) - Suc 0)) \ sint x \ sint x < 2 ^ (LENGTH('a) - Suc 0)\ for x :: \'a::len word\ apply transfer using sbintr_ge sbintr_lt apply auto done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed lemma bintrunc_id: "\m \ int n; 0 < m\ \ take_bit n m = m" by (simp add: take_bit_int_eq_self_iff le_less_trans) lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" if "n = m" "a = c" "take_bit m b = take_bit m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) +lemma bin_sc_pos: + "0 \ i \ 0 \ bin_sc n b i" + by (metis bin_sign_sc sign_Pls_ge_0) + code_identifier code_module Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Word_Lib/Guide.thy b/thys/Word_Lib/Guide.thy --- a/thys/Word_Lib/Guide.thy +++ b/thys/Word_Lib/Guide.thy @@ -1,418 +1,420 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) (*<*) theory Guide imports Word_Lib_Sumo Machine_Word_32 Machine_Word_64 Ancient_Numeral begin context semiring_bit_operations begin lemma bit_eq_iff: \a = b \ (\n. 2 ^ n \ 0 \ bit a n \ bit b n)\ using bit_eq_iff [of a b] by (simp add: possible_bit_def) end notation (output) Generic_set_bit.set_bit (\Generic'_set'_bit.set'_bit\) hide_const (open) Generic_set_bit.set_bit no_notation bit (infixl \!!\ 100) (*>*) section \A short overview over bit operations and word types\ subsection \Key principles\ text \ When formalizing bit operations, it is tempting to represent bit values as explicit lists over a binary type. This however is a bad idea, mainly due to the inherent ambiguities in representation concerning repeating leading bits. Hence this approach avoids such explicit lists altogether following an algebraic path: \<^item> Bit values are represented by numeric types: idealized unbounded bit values can be represented by type \<^typ>\int\, bounded bit values by quotient types over \<^typ>\int\, aka \<^typ>\'a word\. \<^item> (A special case are idealized unbounded bit values ending in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). The fundamental principles are developed in theory \<^theory>\HOL.Bit_Operations\ (which is part of \<^theory>\Main\): \<^item> Multiplication by \<^term>\2 :: int\ is a bit shift to the left and \<^item> Division by \<^term>\2 :: int\ is a bit shift to the right. \<^item> Concerning bounded bit values, iterated shifts to the left may result in eliminating all bits by shifting them all beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. \<^item> The projection on a single bit is then @{thm [mode=iff] bit_iff_odd [where ?'a = int, no_vars]}. \<^item> This leads to the most fundamental properties of bit values: \<^item> Equality rule: @{thm [display, mode=iff] bit_eq_iff [where ?'a = int, no_vars]} \<^item> Induction rule: @{thm [display, mode=iff] bits_induct [where ?'a = int, no_vars]} \<^item> Characteristic properties @{prop [source] \bit (f x) n \ P x n\} are available in fact collection \<^text>\bit_simps\. On top of this, the following generic operations are provided: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ \<^item> Bit mask upto bit \<^term>\n\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> Bitwise negation: @{thm [mode=iff] bit_not_iff_eq [where ?'a = int, no_vars]} \<^item> Bitwise conjunction: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} \<^item> Bitwise disjunction: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} \<^item> Bitwise exclusive disjunction: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} \<^item> Setting a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} \<^item> Unsetting a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} \<^item> Flipping a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm [display] signed_take_bit_def [where ?'a = int, no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} - Bit concatenation on \<^typ>\int\ as given by + Bit concatenation on \<^typ>\int\ as given by @{thm [display] concat_bit_def [no_vars]} appears quite technical but is the logical foundation for the quite natural bit concatenation on \<^typ>\'a word\ (see below). \ subsection \Core word theory\ text \ Proper word types are introduced in theory \<^theory>\HOL-Library.Word\, with the following specific operations: \<^item> Standard arithmetic: @{term \(+) :: 'a::len word \ 'a word \ 'a word\}, @{term \uminus :: 'a::len word \ 'a word\}, @{term \(-) :: 'a::len word \ 'a word \ 'a word\}, @{term \(*) :: 'a::len word \ 'a word \ 'a word\}, @{term \0 :: 'a::len word\}, @{term \1 :: 'a::len word\}, numerals etc. \<^item> Standard bit operations: see above. \<^item> Conversion with unsigned interpretation of words: \<^item> @{term [source] \unsigned :: 'a::len word \ 'b::semiring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \unat :: 'a::len word \ nat\} \<^item> @{term [source] \uint :: 'a::len word \ int\} \<^item> @{term [source] \ucast :: 'a::len word \ 'b::len word\} \<^item> Conversion with signed interpretation of words: \<^item> @{term [source] \signed :: 'a::len word \ 'b::ring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \sint :: 'a::len word \ int\} \<^item> @{term [source] \scast :: 'a::len word \ 'b::len word\} \<^item> Operations with unsigned interpretation of words: \<^item> @{thm [mode=iff] word_le_nat_alt [no_vars]} \<^item> @{thm [mode=iff] word_less_nat_alt [no_vars]} \<^item> @{thm unat_div_distrib [no_vars]} \<^item> @{thm unat_drop_bit_eq [no_vars]} \<^item> @{thm unat_mod_distrib [no_vars]} \<^item> @{thm [mode=iff] udvd_iff_dvd [no_vars]} \<^item> Operations with signed interpretation of words: \<^item> @{thm [mode=iff] word_sle_eq [no_vars]} \<^item> @{thm [mode=iff] word_sless_alt [no_vars]} \<^item> @{thm sint_signed_drop_bit_eq [no_vars]} \<^item> Rotation and reversal: \<^item> @{term [source] \word_rotl :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_rotr :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_roti :: int \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_reverse :: 'a::len word \ 'a word\} \<^item> Concatenation: @{term [source, display] \word_cat :: 'a::len word \ 'b::len word \ 'c::len word\} For proofs about words the following default strategies are applicable: \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_word_eqI\; fact collection \<^text>\bit_simps\). \<^item> Using the @{method transfer} method. \ subsection \More library theories\ text \ Note: currently, most theories listed here are hardly separate entities since they import each other in various ways. Always inspect them to understand what you pull in if you want to import one. \<^descr>[Syntax] \<^descr>[\<^theory>\Word_Lib.Syntax_Bundles\] Bundles to provide alternative syntax for various bit operations. \<^descr>[\<^theory>\Word_Lib.Hex_Words\] Printing word numerals as hexadecimal numerals. \<^descr>[\<^theory>\Word_Lib.Type_Syntax\] Pretty type-sensitive syntax for cast operations. \<^descr>[\<^theory>\Word_Lib.Word_Syntax\] Specific ASCII syntax for prominent bit operations on word. \<^descr>[Proof tools] \<^descr>[\<^theory>\Word_Lib.Norm_Words\] Rewriting word numerals to normal forms. \<^descr>[\<^theory>\Word_Lib.Bitwise\] Method @{method word_bitwise} decomposes word equalities and inequalities into bit propositions. \<^descr>[\<^theory>\Word_Lib.Bitwise_Signed\] Method @{method word_bitwise_signed} decomposes word equalities and inequalities into bit propositions. \<^descr>[\<^theory>\Word_Lib.Word_EqI\] Method @{method word_eqI_solve} decomposes word equalities and inequalities into bit propositions. \<^descr>[Operations] \<^descr>[\<^theory>\Word_Lib.Signed_Division_Word\] Signed division on word: \<^item> @{term [source] \(sdiv) :: 'a::len word \ 'a word \ 'a word\} \<^item> @{term [source] \(smod) :: 'a::len word \ 'a word \ 'a word\} \<^descr>[\<^theory>\Word_Lib.Aligned\] \ \<^item> @{thm [mode=iff] is_aligned_iff_udvd [no_vars]} \<^descr>[\<^theory>\Word_Lib.Least_significant_bit\] The least significant bit as an alias: @{thm [mode=iff] lsb_odd [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Most_significant_bit\] The most significant bit: \<^item> @{thm [mode=iff] msb_int_def [of k]} \<^item> @{thm [mode=iff] word_msb_sint [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_sless_0 [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_bit [no_vars]} \<^descr>[\<^theory>\Word_Lib.Bit_Shifts_Infix_Syntax\] Bit shifts decorated with infix syntax: \<^item> @{thm Bit_Shifts_Infix_Syntax.shiftl_def [no_vars]} \<^item> @{thm Bit_Shifts_Infix_Syntax.shiftr_def [no_vars]} \<^item> @{thm Bit_Shifts_Infix_Syntax.sshiftr_def [no_vars]} \<^descr>[\<^theory>\Word_Lib.Next_and_Prev\] \ \<^item> @{thm word_next_unfold [no_vars]} \<^item> @{thm word_prev_unfold [no_vars]} \<^descr>[\<^theory>\Word_Lib.Enumeration_Word\] More on explicit enumeration of word types. \<^descr>[\<^theory>\Word_Lib.More_Word_Operations\] Even more operations on word. \<^descr>[Types] \<^descr>[\<^theory>\Word_Lib.Signed_Words\] Formal tagging of word types with a \<^text>\signed\ marker. \<^descr>[Lemmas] \<^descr>[\<^theory>\Word_Lib.More_Word\] More lemmas on words. \<^descr>[\<^theory>\Word_Lib.Word_Lemmas\] More lemmas on words, covering many other theories mentioned here. \<^descr>[Words of popular lengths]. \<^descr>[\<^theory>\Word_Lib.Word_8\] for 8-bit words. \<^descr>[\<^theory>\Word_Lib.Word_16\] for 16-bit words. \<^descr>[\<^theory>\Word_Lib.Word_32\] for 32-bit words. \<^descr>[\<^theory>\Word_Lib.Word_64\] - for 64-bit words. + for 64-bit words. This theory is not part of \<^text>\Word_Lib_Sumo\, because it shadows + names from \<^theory>\Word_Lib.Word_32\. They can be used together, but then will have + to use qualified names in applications. \<^descr>[\<^theory>\Word_Lib.Machine_Word_32\ and \<^theory>\Word_Lib.Machine_Word_64\] provide lemmas for 32-bit words and 64-bit words under the same name, which can help to organize applications relying on some form of genericity. \ subsection \More library sessions\ text \ \<^descr>[\<^text>\Native_Word\] Makes machine words and machine arithmetic available for code generation. It provides a common abstraction that hides the differences between the different target languages. The code generator maps these operations to the APIs of the target languages. \ subsection \Legacy theories\ text \ The following theories contain material which has been factored out since it is not recommended to use it in new applications, mostly because matters can be expressed succinctly using already existing operations. This section gives some indication how to migrate away from those theories. However theorem coverage may still be terse in some cases. \<^descr>[\<^theory>\Word_Lib.Word_Lib_Sumo\] An entry point importing any relevant theory in that session. Intended for backward compatibility: start importing this theory when migrating applications to Isabelle2021, and later sort out what you really need. You may need to include \<^theory>\Word_Lib.Word_64\ separately. \<^descr>[\<^theory>\Word_Lib.Generic_set_bit\] Kind of an alias: @{thm set_bit_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Typedef_Morphisms\] A low-level extension to HOL typedef providing conversions along type morphisms. The @{method transfer} method seems to be sufficient for most applications though. \<^descr>[\<^theory>\Word_Lib.Bit_Comprehension\] Comprehension syntax for bit values over predicates \<^typ>\nat \ bool\. For \<^typ>\'a::len word\, straightforward alternatives exist; difficult to handle for \<^typ>\int\. \<^descr>[\<^theory>\Word_Lib.Reversed_Bit_Lists\] Representation of bit values as explicit list in \<^emph>\reversed\ order. This should rarely be necessary: the \<^const>\bit\ projection should be sufficient in most cases. In case explicit lists are needed, existing operations can be used: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Many_More\] Collection of operations and theorems which are kept for backward compatibility and not used in other theories in session \<^text>\Word_Lib\. They are used in applications of \<^text>\Word_Lib\, but should be migrated to there. \ section \Changelog\ text \ \<^descr>[Changes since AFP 2021] ~ - \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is no part of \<^theory>\Word_Lib.Word_Lib_Sumo\ + \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is not part of \<^theory>\Word_Lib.Word_Lib_Sumo\ any longer. \<^item> Infix syntax for \<^term>\(AND)\, \<^term>\(OR)\, \<^term>\(XOR)\ organized in syntax bundle \<^bundle>\bit_operations_syntax\. \<^item> Abbreviation \<^abbrev>\max_word\ moved from distribution into theory \<^theory>\Word_Lib.Legacy_Aliases\. \<^item> Operation \<^const>\test_bit\ replaced by input abbreviation \<^abbrev>\test_bit\. \<^item> Abbreviations \<^abbrev>\bin_nth\, \<^abbrev>\bin_last\, \<^abbrev>\bin_rest\, \<^abbrev>\bintrunc\, \<^abbrev>\sbintrunc\, \<^abbrev>\norm_sint\, \<^abbrev>\bin_cat\ moved into theory \<^theory>\Word_Lib.Legacy_Aliases\. \<^item> Operations \<^abbrev>\bshiftr1\, \<^abbrev>\setBit\, \<^abbrev>\clearBit\ moved from distribution into theory \<^theory>\Word_Lib.Legacy_Aliases\ and replaced by input abbreviations. \<^item> Operations \<^const>\shiftl1\, \<^const>\shiftr1\, \<^const>\sshiftr1\ moved here from distribution. \<^item> Operation \<^const>\complement\ replaced by input abbreviation \<^abbrev>\complement\. \ (*<*) end (*>*) diff --git a/thys/Word_Lib/Machine_Word_32.thy b/thys/Word_Lib/Machine_Word_32.thy --- a/thys/Word_Lib/Machine_Word_32.thy +++ b/thys/Word_Lib/Machine_Word_32.thy @@ -1,136 +1,136 @@ (* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) section "32-Bit Machine Word Setup" theory Machine_Word_32 imports Machine_Word_32_Basics More_Word Bit_Shifts_Infix_Syntax Rsplit begin context includes bit_operations_syntax begin type_synonym machine_word = \machine_word_len word\ lemma word_bits_len_of: \LENGTH(machine_word_len) = word_bits\ - by (simp add: word_bits_conv) + by (simp add: word_bits_conv) lemma word_bits_size: "size (w :: machine_word) = word_bits" by (simp add: word_bits_def word_size) lemma word_bits_word_size_conv: \word_bits = word_size * 8\ by (simp add: word_bits_def word_size_def) lemma word_size_word_size_bits: \word_size = (2 :: 'a :: semiring_1) ^ word_size_bits\ by (simp add: word_size_def word_size_bits_def) lemma lt_word_bits_lt_pow: "sz < word_bits \ sz < 2 ^ word_bits" by (simp add: word_bits_conv) lemma if_then_1_else_0: "((if P then 1 else 0) = (0 :: machine_word)) = (\ P)" by simp lemma if_then_0_else_1: "((if P then 0 else 1) = (0 :: machine_word)) = (P)" by simp lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 lemma bool_mask [simp]: \0 < x AND 1 \ x AND 1 = 1\ for x :: machine_word by (rule bool_mask') auto lemma in_16_range: "0 \ S \ r \ (\x. r + x * (16 :: machine_word)) ` S" "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: machine_word. r + x * 16) ` S" by (clarsimp simp: image_def elim!: bexI[rotated])+ lemma le_step_down_word_3: fixes x :: machine_word shows "\x \ y; x \ y\ \ x \ y - 1" by (fact le_step_down_word_2) lemma shiftr_1: "(x::machine_word) >> 1 = 0 \ x < 2" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_Suc) done lemma Suc_unat_mask_div: "Suc (unat (mask sz div word_size :: machine_word)) = 2 ^ (min sz word_bits - word_size_bits)" by (simp add: word_size_word_size_bits unat_drop_bit_eq unat_mask_eq drop_bit_mask_eq Suc_mask_eq_exp flip: drop_bit_eq_div word_bits_conv) lemma ucast_not_helper: fixes a::"8 word" assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::machine_word)" proof assume "ucast a = (0xFF::machine_word)" also have "(0xFF::machine_word) = ucast (0xFF::8 word)" by simp finally show False using a apply - apply (drule up_ucast_inj, simp) apply simp done qed lemma unat_less_2p_word_bits: "unat (x :: machine_word) < 2 ^ word_bits" apply (simp only: word_bits_def) apply (rule unat_lt2p) done lemma unat_less_word_bits: fixes y :: machine_word shows "x < unat y \ x < 2 ^ word_bits" unfolding word_bits_def by (rule order_less_trans [OF _ unat_lt2p]) lemma unat_mask_2_less_4: "unat (p AND mask 2 :: machine_word) < 4" by (rule unat_less_helper) (simp only: take_bit_eq_mod word_mod_less_divisor flip: take_bit_eq_mask, simp add: word_mod_less_divisor) lemma unat_mult_simple: \unat (x * y) = unat x * unat y\ if \unat x * unat y < 2 ^ LENGTH(machine_word_len)\ for x y :: machine_word using that by (simp flip: unat_mult_lem) lemma upto_2_helper: "{0..<2 :: machine_word} = {0, 1}" by (safe; simp) unat_arith lemma word_ge_min: \- (2 ^ (word_bits - 1)) \ sint x\ for x :: machine_word using sint_ge [of x] by (simp add: word_bits_def) lemma word_rsplit_0: "word_rsplit (0 :: machine_word) = replicate (word_bits div 8) (0 :: 8 word)" by (simp add: word_rsplit_def bin_rsplit_def word_bits_def word_size_def Cons_replicate_eq) lemma x_less_2_0_1: fixes x :: machine_word shows "x < 2 \ x = 0 \ x = 1" by (rule x_less_2_0_1') auto end end diff --git a/thys/Word_Lib/Most_significant_bit.thy b/thys/Word_Lib/Most_significant_bit.thy --- a/thys/Word_Lib/Most_significant_bit.thy +++ b/thys/Word_Lib/Most_significant_bit.thy @@ -1,194 +1,194 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Dedicated operation for the most significant bit\ theory Most_significant_bit imports "HOL-Library.Word" Bit_Shifts_Infix_Syntax More_Word More_Arithmetic begin class msb = fixes msb :: \'a \ bool\ instantiation int :: msb begin definition \msb x \ x < 0\ for x :: int instance .. end lemma msb_bin_rest [simp]: "msb (x div 2) = msb x" for x :: int by (simp add: msb_int_def) context includes bit_operations_syntax begin lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" by(simp add: msb_int_def not_less) end lemma msb_shiftl [simp]: "msb ((x :: int) << n) \ msb x" by (simp add: msb_int_def shiftl_def) lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \ msb x" by (simp add: msb_int_def shiftr_def) lemma msb_0 [simp]: "msb (0 :: int) = False" by(simp add: msb_int_def) lemma msb_1 [simp]: "msb (1 :: int) = False" by(simp add: msb_int_def) lemma msb_numeral [simp]: "msb (numeral n :: int) = False" "msb (- numeral n :: int) = True" by(simp_all add: msb_int_def) instantiation word :: (len) msb begin definition msb_word :: \'a word \ bool\ where msb_word_iff_bit: \msb w \ bit w (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ instance .. end lemma msb_word_eq: \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ by (simp add: msb_word_iff_bit) lemma word_msb_sint: "msb w \ sint w < 0" by (simp add: msb_word_eq bit_last_iff) lemma msb_word_iff_sless_0: \msb w \ w by (simp add: word_msb_sint word_sless_alt) lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bit x (LENGTH('a) - 1)" by (simp add: msb_word_iff_bit bit_simps) lemma word_msb_numeral [simp]: "msb (numeral w::'a::len word) = bit (numeral w :: int) (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule msb_word_of_int) lemma word_msb_neg_numeral [simp]: "msb (- numeral w::'a::len word) = bit (- numeral w :: int) (LENGTH('a) - 1)" unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" by (simp add: msb_word_iff_bit) lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" by (simp add: msb_word_iff_bit le_Suc_eq) lemma word_msb_nth: "msb w = bit (uint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: msb_word_iff_bit bit_simps) lemma msb_nth: "msb w = bit w (LENGTH('a) - 1)" for w :: "'a::len word" by (fact msb_word_eq) - + lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" by (simp add: msb_word_eq not_le) lemma msb_shift: "msb w \ w >> LENGTH('a) - 1 \ 0" for w :: "'a::len word" by (simp add: drop_bit_eq_zero_iff_not_bit_last msb_word_eq shiftr_def) lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_iff_bit word_size) apply transfer apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" apply (simp add: word_sle_eq word_sint_msb_eq word_size word_le_def) apply safe apply (rule order_trans[OF _ uint_ge_0]) apply (simp add: order_less_imp_le) apply (erule notE[OF leD]) apply (rule order_less_le_trans[OF _ uint_ge_0]) apply simp done lemma word_sless_msb_less: "x (msb y \ msb x) \ ((msb x \ \ msb y) \ x < y)" by (auto simp add: word_sless_eq word_sle_msb_le) lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply (simp add: bit_simps) done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_iff_bit) apply transfer apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" apply (cases \LENGTH('a)\) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) apply (metis le_less_Suc_eq test_bit_bin) done lemma msb_ucast_eq: "LENGTH('a) = LENGTH('b) \ msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" by (simp add: msb_word_eq bit_simps) lemma msb_big: \msb a \ 2 ^ (LENGTH('a) - Suc 0) \ a\ for a :: \'a::len word\ using bang_is_le [of a \LENGTH('a) - Suc 0\] apply (auto simp add: msb_nth word_le_not_less) apply (rule ccontr) apply (erule notE) apply (rule ccontr) apply (clarsimp simp: not_less) apply (subgoal_tac "a = take_bit (LENGTH('a) - Suc 0) a") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply auto apply (simp flip: take_bit_eq_mask) apply (rule sym) apply (simp add: take_bit_eq_self_iff_drop_bit_eq_0 drop_bit_eq_zero_iff_not_bit_last) done end diff --git a/thys/Word_Lib/Signed_Words.thy b/thys/Word_Lib/Signed_Words.thy --- a/thys/Word_Lib/Signed_Words.thy +++ b/thys/Word_Lib/Signed_Words.thy @@ -1,113 +1,113 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Signed Words" theory Signed_Words imports "HOL-Library.Word" begin text \Signed words as separate (isomorphic) word length class. Useful for tagging words in C.\ typedef ('a::len0) signed = "UNIV :: 'a set" .. lemma card_signed [simp]: "CARD (('a::len0) signed) = CARD('a)" unfolding type_definition.card [OF type_definition_signed] by simp instantiation signed :: (len0) len0 begin definition len_signed [simp]: "len_of (x::'a::len0 signed itself) = LENGTH('a)" instance .. end instance signed :: (len) len by (intro_classes, simp) lemma scast_scast_id [simp]: "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" by (auto simp: is_up scast_up_scast_id) lemma ucast_scast_id [simp]: "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" by transfer (simp add: take_bit_signed_take_bit) lemma scast_of_nat [simp]: "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" by transfer (simp add: take_bit_signed_take_bit) lemma scast_ucast_id [simp]: "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" - by transfer (simp add: take_bit_signed_take_bit) + by transfer (simp add: take_bit_signed_take_bit) lemma scast_eq_scast_id [simp]: "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" by (metis ucast_scast_id) lemma ucast_eq_ucast_id [simp]: "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" by (metis scast_ucast_id) lemma scast_ucast_norm [simp]: "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" by (metis scast_ucast_id ucast_scast_id)+ lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (fact of_nat_unat) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" by transfer (simp add: bit_simps) lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ (ucast x :: ('b :: len) signed word) 'a::len word\ apply (cases \LENGTH('b)\) apply simp_all apply transfer apply (simp add: signed_take_bit_take_bit) - apply (metis add.commute mask_eq_exp_minus_1 mask_eq_take_bit_minus_one take_bit_incr_eq zle_add1_eq_le) + apply (metis add.commute mask_eq_exp_minus_1 take_bit_incr_eq zle_add1_eq_le) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply transfer apply (cases \LENGTH('a)\) apply (simp_all add: take_bit_Suc_from_most bit_simps) apply (simp_all add: bit_simps disjunctive_add) done lemma nth_w2p_scast: "(bit (scast ((2::'a::len signed word) ^ n) :: 'a word) m) \ (bit (((2::'a::len word) ^ n) :: 'a word) m)" by (simp add: bit_simps) lemma scast_nop1 [simp]: "((scast ((of_int x)::('a::len) word))::'a signed word) = of_int x" - by transfer (simp add: take_bit_signed_take_bit) + by transfer (simp add: take_bit_signed_take_bit) lemma scast_nop2 [simp]: "((scast ((of_int x)::('a::len) signed word))::'a word) = of_int x" by transfer (simp add: take_bit_signed_take_bit) lemmas scast_nop = scast_nop1 scast_nop2 scast_id type_synonym 'a sword = "'a signed word" end diff --git a/thys/Word_Lib/Word_EqI.thy b/thys/Word_Lib/Word_EqI.thy --- a/thys/Word_Lib/Word_EqI.thy +++ b/thys/Word_Lib/Word_EqI.thy @@ -1,72 +1,73 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Solving Word Equalities" theory Word_EqI imports More_Word Aligned "HOL-Eisbach.Eisbach_Tools" begin text \ Some word equalities can be solved by considering the problem bitwise for all @{prop "n < LENGTH('a::len)"}, which is different to running @{text word_bitwise} and expanding into an explicit list of bits. \ named_theorems word_eqI_simps lemmas [word_eqI_simps] = word_ops_nth_size + word_ao_nth bit_mask_iff word_size word_or_zero neg_mask_test_bit nth_ucast nth_w2p bit_push_bit_iff bit_drop_bit_eq less_2p_is_upper_bits_unset le_mask_high_bits bang_eq neg_test_bit is_up is_down is_aligned_nth neg_mask_le_high_bits lemmas word_eqI_rule = word_eqI [rule_format] lemma test_bit_lenD: "bit x n \ n < LENGTH('a) \ bit x n" for x :: "'a :: len word" by (fastforce dest: test_bit_size simp: word_size) method word_eqI uses simp simp_del split split_del cong flip = ((* reduce conclusion to test_bit: *) rule word_eqI_rule, (* make sure we're in clarsimp normal form: *) (clarsimp simp: simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* turn x < 2^n assumptions into mask equations: *) ((drule less_mask_eq)+)?, (* expand and distribute test_bit everywhere: *) (clarsimp simp: word_eqI_simps simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* add any additional word size constraints to new indices: *) ((drule test_bit_lenD)+)?, (* try to make progress (can't use +, would loop): *) (clarsimp simp: word_eqI_simps simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* helps sometimes, rarely: *) (simp add: simp test_bit_conj_lt del: simp_del flip: flip split: split split del: split_del cong: cong)?) method word_eqI_solve uses simp simp_del split split_del cong flip = solves \word_eqI simp: simp simp_del: simp_del split: split split_del: split_del cong: cong simp flip: flip; (fastforce dest: test_bit_size simp: word_eqI_simps simp flip: flip simp: simp simp del: simp_del split: split split del: split_del cong: cong)?\ end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,1961 +1,1970 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Signed_Division_Word Signed_Words More_Word Most_significant_bit Enumeration_Word Aligned Bit_Shifts_Infix_Syntax + Word_EqI begin context includes bit_operations_syntax begin lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma ucast_zero_is_aligned: \is_aligned w n\ if \UCAST('a::len \ 'b::len) w = 0\ \n \ LENGTH('b)\ proof (rule is_aligned_bitI) fix q assume \q < n\ moreover have \bit (UCAST('a::len \ 'b::len) w) q = bit 0 q\ using that by simp with \q < n\ \n \ LENGTH('b)\ show \\ bit w q\ by (simp add: bit_simps) qed lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma le_max_word_ucast_id: \UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by (simp add: of_int_mask_eq) have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis apply (simp add: unsigned_ucast_eq take_bit_word_eq_self_iff) apply (meson \x \ 2 ^ LENGTH('b) - 1\ not_le word_less_sub_le) done qed lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ by (rule bit_eqI) (simp flip: drop_bit_eq_div add: bit_simps) lemma bit_shiftl_word_iff [bit_simps]: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: bit_simps) lemma bit_shiftr_word_iff: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: bit_simps) lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by (rule bit_eqI) (auto simp add: bit_simps not_less simp flip: drop_bit_eq_div dest: bit_imp_le_length) lemma sshiftr_0: "0 >>> n = 0" by (simp add: sshiftr_def) - + lemma sshiftr_n1: "-1 >>> n = -1" by (simp add: sshiftr_def) lemma bit_sshiftr_word_iff: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ by (fact bit_sshiftr_iff) lemma nth_sshiftr : "bit (w >>> m) n = (n < size w \ (if n + m \ size w then bit w (size w - 1) else bit w (n + m)))" apply (auto simp add: bit_simps word_size ac_simps not_less) apply (meson bit_imp_le_length bit_shiftr_word_iff leD) done lemma sshiftr_numeral: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral k) >> numeral n)\ using signed_drop_bit_word_numeral [of n k] by (simp add: sshiftr_def shiftr_def) lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" apply (rule bit_eqI) apply (cases \n < LENGTH('a)\) apply (auto simp add: bit_simps not_less le_diff_conv2 simp flip: drop_bit_eq_div) done lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by (simp add: mask_eq_exp_minus_1 shiftl_def) lemma nth_shiftl': "bit (w << m) n \ n < size w \ n >= m \ bit w (n - m)" for w :: "'a::len word" by (simp add: bit_simps word_size ac_simps) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr: "bit (w >> m) n = bit w (n + m)" for w :: "'a::len word" by (simp add: bit_simps ac_simps) lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" by (fact uint_shiftr_eq) lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" by (simp add: shiftl_def word_size) lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (simp add: shiftl_def push_bit_eq_mult) lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) +lemma word_shift_by_3: + "x * 8 = (x::'a::len word) << 3" + by (simp add: shiftl_t2n) + lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) lemma shiftr_x_0 [simp]: "x >> 0 = x" for x :: "'a::len word" by (simp add: shiftr_def) lemma shiftl_x_0 [simp]: "x << 0 = x" for x :: "'a::len word" by (simp add: shiftl_def) lemma shiftl_1: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_def) lemma shiftr_1 [simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (simp add: shiftr_def) lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add: uint_2p_alt word_size) apply (metis uint_shiftr_eq word_of_int_uint) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply (unfold shiftr_def) apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_eq_div zdiv_mono1) done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (metis le_cases le_shiftr verit_la_disequality) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size bit_simps) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ by (simp flip: take_bit_eq_mask add: shiftr_def take_bit_eq_self_iff_drop_bit_eq_0) lemma mask_shiftl_decompose: "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" by (unfold shiftl_def) (fact push_bit_and) lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" by (unfold shiftr_def) (fact drop_bit_and) lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" by (unfold shiftl_def) (fact push_bit_or) lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" by (unfold shiftr_def) (fact drop_bit_or) lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" by (rule bit_word_eqI) (simp add: bit_simps) lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth bit_simps) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma mask_shift: "(x AND NOT (mask y)) >> y = x >> y" for x :: \'a::len word\ apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_eq_unatI) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by (transfer fixing: n) (simp add: take_bit_push_bit) lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') apply transfer apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) done lemma shiftr_less_t2n': "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit ac_simps) done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (rule bit_word_eqI) (simp add: bit_simps) lemma signed_shift_guard_to_word: \unat x * 2 ^ y < 2 ^ n \ x = 0 \ x < 1 << n >> y\ if \n < LENGTH('a)\ \0 < n\ for x :: \'a::len word\ proof (cases \x = 0\) case True then show ?thesis by simp next case False then have \unat x \ 0\ by (simp add: unat_eq_0) then have \unat x \ 1\ by simp show ?thesis proof (cases \y < n\) case False then have \n \ y\ by simp then obtain q where \y = n + q\ using le_Suc_ex by blast moreover have \(2 :: nat) ^ n >> n + q \ 1\ by (simp add: drop_bit_eq_div power_add shiftr_def) ultimately show ?thesis using \x \ 0\ \unat x \ 1\ \n < LENGTH('a)\ by (simp add: power_add not_less word_le_nat_alt unat_drop_bit_eq shiftr_def shiftl_def) next case True with that have \y < LENGTH('a)\ by simp show ?thesis proof (cases \2 ^ n = unat x * 2 ^ y\) case True moreover have \unat x * 2 ^ y < 2 ^ LENGTH('a)\ using \n < LENGTH('a)\ by (simp flip: True) moreover have \(word_of_nat (2 ^ n) :: 'a word) = word_of_nat (unat x * 2 ^ y)\ using True by simp then have \2 ^ n = x * 2 ^ y\ by simp ultimately show ?thesis using \y < LENGTH('a)\ by (auto simp add: drop_bit_eq_div word_less_nat_alt unat_div unat_word_ariths shiftr_def shiftl_def) next case False with \y < n\ have *: \unat x \ 2 ^ n div 2 ^ y\ by (auto simp flip: power_sub power_add) have \unat x * 2 ^ y < 2 ^ n \ unat x * 2 ^ y \ 2 ^ n\ using False by (simp add: less_le) also have \\ \ unat x \ 2 ^ n div 2 ^ y\ by (simp add: less_eq_div_iff_mult_less_eq) also have \\ \ unat x < 2 ^ n div 2 ^ y\ using * by (simp add: less_le) finally show ?thesis using that \x \ 0\ by (simp flip: push_bit_eq_mult drop_bit_eq_div add: shiftr_def shiftl_def unat_drop_bit_eq word_less_iff_unsigned [where ?'a = nat]) qed qed qed lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" by (rule bit_word_eqI) (auto simp add: bit_simps word_size dest: bit_imp_le_length) lemma shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask add: take_bit_push_bit shiftl_def) lemma rshift_sub_mask_eq: "(a >> (size a - b)) AND mask b = a >> (size a - b)" for a :: \'a::len word\ using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" for a :: \'a::len word\ apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" for w :: \'a::len word\ by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma word_and_1_shiftl: "x AND (1 << n) = (if bit x n then (1 << n) else 0)" for x :: "'a :: len word" apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x AND (mask n << m) = ((x >> m) AND mask n) << m" for x :: \'a::len word\ apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma of_bool_nth: "of_bool (bit x v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit word_and_1 shiftr_def) lemma shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit) done lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (bit x n))" by (simp add: and_exp_eq_0_iff_not_bit shiftl_def) lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" for x :: \'a::len word\ by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" for a b :: \'a::len word\ apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma shiftr1_unfold: "x div 2 = x >> 1" by (simp add: drop_bit_eq_div shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by (simp add: drop_bit_eq_div shiftr_def) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma shiftr1_irrelevant_lsb: "bit (x::('a::len) word) 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb': "\ (bit (x::('a::len) word) 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) OR (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) (auto simp add: word_size bit_not_iff bit_push_bit_iff bit_mask_iff shiftl_def) (* Comparisons between different word sizes. *) lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (auto simp add: take_bit_word_eq_self_iff word_less_nat_alt shiftr_def simp flip: take_bit_eq_self_iff_drop_bit_eq_0 intro: ccontr) lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) (* continue sorting out from here *) (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a AND b = 0" by (metis and_zero_eq is_aligned_mask le_mask_imp_and_mask word_bw_lcs(1)) lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a OR b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a OR b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (simp add: bit_ucast_iff is_aligned_nth) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" apply (simp add: less_eq_mask_iff_take_bit_eq_self) apply transfer apply (simp add: ac_simps) done lemma shiftl_inj: \x = y\ if \x << n = y << n\ \x \ mask (LENGTH('a) - n)\ \y \ mask (LENGTH('a) - n)\ for x y :: \'a::len word\ proof (cases \n < LENGTH('a)\) case False with that show ?thesis by simp next case True moreover from that have \take_bit (LENGTH('a) - n) x = x\ \take_bit (LENGTH('a) - n) y = y\ by (simp_all add: less_eq_mask_iff_take_bit_eq_self) ultimately show ?thesis using \x << n = y << n\ by (metis diff_less gr_implies_not0 linorder_cases linorder_not_le shiftl_shiftr_id shiftl_x_0 take_bit_word_eq_self_iff) qed lemma distinct_word_add_ucast_shift_inj: \p' = p \ off' = off\ if *: \p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n)\ and \is_aligned p n'\ \is_aligned p' n'\ \n' = n + LENGTH('a)\ \n' < LENGTH('b)\ proof - from \n' = n + LENGTH('a)\ have [simp]: \n' - n = LENGTH('a)\ \n + LENGTH('a) = n'\ by simp_all from \is_aligned p n'\ obtain q where p: \p = push_bit n' (word_of_nat q)\ \q < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') from \is_aligned p' n'\ obtain q' where p': \p' = push_bit n' (word_of_nat q')\ \q' < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') define m :: nat where \m = unat off\ then have off: \off = word_of_nat m\ by simp define m' :: nat where \m' = unat off'\ then have off': \off' = word_of_nat m'\ by simp have \push_bit n' q + take_bit n' (push_bit n m) < 2 ^ LENGTH('b)\ by (metis id_apply is_aligned_no_wrap''' of_nat_eq_id of_nat_push_bit p(1) p(2) take_bit_nat_eq_self_iff take_bit_nat_less_exp take_bit_push_bit that(2) that(5) unsigned_of_nat) moreover have \push_bit n' q' + take_bit n' (push_bit n m') < 2 ^ LENGTH('b)\ by (metis \n' - n = LENGTH('a)\ id_apply is_aligned_no_wrap''' m'_def of_nat_eq_id of_nat_push_bit off' p'(1) p'(2) take_bit_nat_eq_self_iff take_bit_push_bit that(3) that(5) unsigned_of_nat) ultimately have \push_bit n' q + take_bit n' (push_bit n m) = push_bit n' q' + take_bit n' (push_bit n m')\ using * by (simp add: p p' off off' push_bit_of_nat push_bit_take_bit word_of_nat_inj unsigned_of_nat shiftl_def flip: of_nat_add) then have \int (push_bit n' q + take_bit n' (push_bit n m)) = int (push_bit n' q' + take_bit n' (push_bit n m'))\ by simp then have \concat_bit n' (int (push_bit n m)) (int q) = concat_bit n' (int (push_bit n m')) (int q')\ by (simp add: of_nat_push_bit of_nat_take_bit concat_bit_eq) then show ?thesis by (simp add: p p' off off' take_bit_of_nat take_bit_push_bit word_of_nat_eq_iff concat_bit_eq_iff) (simp add: push_bit_eq_mult) qed lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ by (metis (no_types, opaque_lifting) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: of_nat_diff unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y AND mask n = 0 \ x + y AND NOT(mask n) = (x AND NOT(mask n)) + y" for x y :: \'a::len word\ by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by (rule bit_word_eqI) (auto simp add: bit_simps dest: bit_imp_le_length) lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x AND mask n = 0; y AND mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y AND mask n = mask n \ (x AND y) AND mask n = x AND mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y AND mask n = 0 \ (x AND y) AND mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) AND b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) AND b < c" by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ by (fact unsigned_minus_1_eq_mask) lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by (drule sym) (simp flip: take_bit_eq_mask add: unsigned_ucast_eq) lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by (simp add: word_eq_iff bit_simps) lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x AND mask n = 0; m \ n \ \ x AND mask m = 0" for x :: \'a::len word\ by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) AND mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) AND 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" apply (drule sym) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ bit p n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr ac_simps) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" by (simp add: word_le_nat_alt unsigned_of_nat take_bit_nat_eq_self) lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x AND NOT(mask sz)) + 2 ^ sz - 1" for x :: \'a::len word\ by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, opaque_lifting) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) AND (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) AND mask m) = unat (x AND mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit) apply (rule order_trans) defer apply assumption apply (simp add: nat_le_iff of_nat_diff) done lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n AND msk = mask n \ \ (x mod m) AND msk = x mod m" for x :: \'a::len word\ by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x AND mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m AND NOT(mask n :: 'a::len word)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr AND NOT(mask sz')) - (ptr AND NOT(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") for ptr :: \'a::len word\ proof - assume lt: "sz' \ sz" hence "?lhs = ptr AND (mask sz AND NOT(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) AND NOT(mask sz) = 0" by (simp add: of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr AND NOT(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) AND NOT(mask sz)) + unat (ptr AND mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: \unat (x << n) < 2 ^ m\ if \unat (x :: 'a :: len word) < 2 ^ (m - n)\ \m < LENGTH('a)\ proof (cases \n \ m\) case False with that show ?thesis apply (transfer fixing: m n) apply (simp add: not_le take_bit_push_bit) - apply (metis diff_le_self order_le_less_trans push_bit_of_0 take_bit_0 take_bit_int_eq_self - take_bit_int_less_exp take_bit_nonnegative take_bit_tightened) + apply (metis diff_le_self order_le_less_trans push_bit_of_0 take_bit_0 take_bit_int_eq_self + take_bit_int_less_exp take_bit_nonnegative take_bit_tightened) done next case True moreover define q r where \q = m - n\ and \r = LENGTH('a) - n - q\ ultimately have \m - n = q\ \m = n + q\ \LENGTH('a) = r + q + n\ using that by simp_all with that show ?thesis apply (transfer fixing: m n q r) apply (simp add: not_le take_bit_push_bit) apply (simp add: push_bit_eq_mult power_add) using take_bit_tightened_less_eq_int [of \r + q\ \r + q + n\] apply (rule le_less_trans) apply simp_all done qed lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d AND mask n) = unat d \ unat (p + d AND NOT(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) AND NOT(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) AND mask high_bits = x >> low_bits" for x :: \'a::len word\ by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x AND mask low_bits = 0) = (x = 0)" for x :: \'a::len word\ using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) AND mask n = q AND mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) AND mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x AND NOT(mask sz)) + (x AND mask sz) \ (x AND NOT(mask sz)) + ((x AND mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "NOT(mask n :: 'a::len word) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * NOT(mask n) = - (x AND NOT(mask n))" for x :: \'a::len word\ by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) AND NOT(mask n) = p + (x AND NOT(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) AND NOT(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) +lemma plus_mask_AND_NOT_mask_eq: + "x AND NOT(mask n) = x \ (x + mask n) AND NOT(mask n) = x" for x::\'a::len word\ + by (subst word_plus_and_or_coroll; word_eqI_solve) + lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" by (rule word_eqI) (auto simp add: bit_simps) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p AND mask a >> b) :: 'a :: len word) = p AND mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p AND mask a >> b) :: 'a :: len word) = unat (p AND mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p AND mask a) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p AND mask a << c) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p AND NOT(mask a)) AND mask b = 0" for p :: \'a::len word\ using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p AND NOT(mask a)) + ((p AND mask a) AND NOT(mask b)) + (p AND mask b) = p" for p :: \'a::len word\ by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p AND mask a >> b << b) = (p AND mask a) AND NOT(mask b)" for p :: \'a::len word\ by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" apply (simp add: shiftl_def shiftr_def flip: push_bit_eq_mult take_bit_eq_mask word_unat_eq_iff) apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (rule bit_word_eqI) (simp add: bit_simps) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x AND mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (smt (z3) eq_or_less_helperD le_add2 le_eq_less_or_eq le_trans power_add unat_mult_lem unat_pow_le_intro unat_power_lower word_eq_unatI) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w AND mask (size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" - by (simp add: shiftl_def minus_exp_eq_not_mask push_bit_minus_one_eq_not_mask) + by (simp add: shiftl_def minus_exp_eq_not_mask) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" by transfer (simp flip: take_bit_eq_mask add: ac_simps) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "signed_take_bit n (uint (ucast w :: 'b word)) = signed_take_bit n (uint w)" by (rule bit_eqI) (use assms in \simp add: bit_simps\) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit n (uint w)) :: 'a word) i = (if n < i then bit w n else bit w i)" using assms by (simp add: bit_simps) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) i = (if LENGTH('b::len) \ i then bit w (LENGTH('b) - 1) else bit w i)" using len_a by (auto simp add: sbintrunc_uint_ucast bit_simps) lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. bit w i = bit w (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ NOT(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" apply (simp flip: take_bit_eq_mask) apply transfer apply simp done lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" by (rule bit_word_eqI) (auto simp add: bit_simps is_down.rep_eq) lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (metis local.distrib local.is_down take_bit_eq_mod ucast_down_wi uint_word_of_int_eq word_of_int_uint) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) let ?range = \{- (2 ^ (size a - 1))..<2 ^ (size a - 1)} :: int set\ have result_range: "sint a sdiv sint b \ ?range \ {2 ^ (size a - 1)}" using sdiv_word_min [of a b] sdiv_word_max [of a b] by auto have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: signed_divide_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply auto apply (cases \size a\) apply simp_all apply (smt (z3) One_nat_def diff_Suc_1 signed_word_eqI sint_int_min sint_range_size wsst_TYs(3)) done have result_range_simple: "(sint a sdiv sint b \ ?range) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith signed_take_bit_int_eq_self_iff intro: sym dest: sym) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ lemma nasty_split_lt: \x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1\ if \x < 2 ^ (m - n)\ \n \ m\ \m < LENGTH('a::len)\ for x :: \'a::len word\ proof - define q where \q = m - n\ with \n \ m\ have \m = q + n\ by simp with \x < 2 ^ (m - n)\ have *: \i < q\ if \bit x i\ for i using that by simp (metis bit_take_bit_iff take_bit_word_eq_self_iff) from \m = q + n\ have \push_bit n x OR mask n \ mask m\ by (auto simp add: le_mask_high_bits word_size bit_simps dest!: *) then have \push_bit n x + mask n \ mask m\ by (simp add: disjunctive_add bit_simps) then show ?thesis by (simp add: mask_eq_exp_minus_1 push_bit_eq_mult) qed lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done end end diff --git a/thys/Word_Lib/Word_Lib_Sumo.thy b/thys/Word_Lib/Word_Lib_Sumo.thy --- a/thys/Word_Lib/Word_Lib_Sumo.thy +++ b/thys/Word_Lib/Word_Lib_Sumo.thy @@ -1,137 +1,136 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) section \Ancient comprehensive Word Library\ theory Word_Lib_Sumo imports "HOL-Library.Word" Aligned Bit_Comprehension Bit_Shifts_Infix_Syntax Bits_Int Bitwise_Signed Bitwise Enumeration_Word Generic_set_bit Hex_Words Least_significant_bit More_Arithmetic More_Divides More_Sublist Even_More_List More_Misc Strict_part_mono Legacy_Aliases Most_significant_bit Next_and_Prev Norm_Words Reversed_Bit_Lists Rsplit Signed_Words Syntax_Bundles Typedef_Morphisms Type_Syntax Word_EqI Word_Lemmas Word_8 Word_16 Word_32 - Word_64 Word_Syntax Signed_Division_Word Singleton_Bit_Shifts More_Word_Operations Many_More begin unbundle bit_operations_syntax unbundle bit_projection_infix_syntax declare word_induct2[induct type] declare word_nat_cases[cases type] declare signed_take_bit_Suc [simp] (* these generate take_bit terms, which we often don't want for concrete lengths *) lemmas of_int_and_nat = unsigned_of_nat unsigned_of_int signed_of_int signed_of_nat bundle no_take_bit begin declare of_int_and_nat[simp del] end lemmas bshiftr1_def = bshiftr1_eq lemmas is_down_def = is_down_eq lemmas is_up_def = is_up_eq lemmas mask_def = mask_eq lemmas scast_def = scast_eq lemmas shiftl1_def = shiftl1_eq lemmas shiftr1_def = shiftr1_eq lemmas sshiftr1_def = sshiftr1_eq lemmas sshiftr_def = sshiftr_eq_funpow_sshiftr1 lemmas to_bl_def = to_bl_eq lemmas ucast_def = ucast_eq lemmas unat_def = unat_eq_nat_uint lemmas word_cat_def = word_cat_eq lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl lemmas word_rotl_def = word_rotl_eq lemmas word_rotr_def = word_rotr_eq lemmas word_sle_def = word_sle_eq lemmas word_sless_def = word_sless_eq lemmas uint_0 = uint_nonnegative lemmas uint_lt = uint_bounded lemmas uint_mod_same = uint_idem lemmas of_nth_def = word_set_bits_def lemmas of_nat_word_eq_iff = word_of_nat_eq_iff lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff lemmas of_int_word_eq_iff = word_of_int_eq_iff lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff lemmas word_next_def = word_next_unfold lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp word32_and_max_simp lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemmas and_bang = word_and_nth lemmas sdiv_int_def = signed_divide_int_def lemmas smod_int_def = signed_modulo_int_def (* shortcut for some specific lengths *) lemma word_fixed_sint_1[simp]: "sint (1::8 word) = 1" "sint (1::16 word) = 1" "sint (1::32 word) = 1" "sint (1::64 word) = 1" by (auto simp: sint_word_ariths) declare of_nat_diff [simp] (* Haskellish names/syntax *) notation (input) bit ("testBit") lemmas cast_simps = cast_simps ucast_down_bl (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by (auto simp add: bit_simps not_le dest: bit_imp_le_length) end diff --git a/web/about.html b/web/about.html --- a/web/about.html +++ b/web/about.html @@ -1,146 +1,146 @@ Archive of Formal Proofs

 

 

 

 

 

 

Archive of Formal Proofs

 

About

The Archive of Formal Proofs is a collection of proof libraries, examples, and larger scientific developments, mechanically checked in the theorem prover Isabelle. It is organized in the way of a scientific journal. Submissions are refereed.

The archive repository is hosted on Heptapod to provide easy free access to archive entries. The entries are tested and maintained continuously against the current stable release of Isabelle. Older versions of archive entries will remain available.

Editors

The editors of the archive are

Why

We aim to strengthen the community and to foster the development of formal proofs.

We hope that the archive will provide

  • a resource of knowledge, examples, and libraries for users,
  • a large and relevant test bed of theories for Isabelle developers, and
  • a central, citable place for authors to publish their theories

We encourage authors of publications that contain Isabelle developments to make their theories available in the Archive of Formal Proofs and to refer to the archive entry in their publication. It makes it easier for referees to check the validity of theorems (all entries in the archive are mechanically checked), it makes it easier for readers of the publication to understand details of your development, and it makes it easier to use and build on your work.

License

All entries in the Archive of Formal Proofs are licensed under a BSD-style License or the GNU LGPL. This means they are free to download, free to use, free to change, and free to redistribute with minimal restrictions.

The authors retain their full copyright on their original work, including their right to make the development available under another, additional license in the future.

\ No newline at end of file diff --git a/web/entries/Amicable_Numbers.html b/web/entries/Amicable_Numbers.html --- a/web/entries/Amicable_Numbers.html +++ b/web/entries/Amicable_Numbers.html @@ -1,192 +1,192 @@ Amicable Numbers - Archive of Formal Proofs

 

 

 

 

 

 

Amicable Numbers

 

Title: Amicable Numbers
Author: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki
Submission date: 2020-08-04
Abstract: This is a formalisation of Amicable Numbers, involving some relevant material including Euler's sigma function, some relevant definitions, results and examples as well as rules such as Thābit ibn Qurra's Rule, Euler's Rule, te Riele's Rule and Borho's Rule with breeders.
BibTeX:
@article{Amicable_Numbers-AFP,
   author  = {Angeliki Koutsoukou-Argyraki},
   title   = {Amicable Numbers},
   journal = {Archive of Formal Proofs},
   month   = aug,
   year    = 2020,
   note    = {\url{https://isa-afp.org/entries/Amicable_Numbers.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on: Polynomial_Factorization, Pratt_Certificate

\ No newline at end of file diff --git a/web/entries/Aristotles_Assertoric_Syllogistic.html b/web/entries/Aristotles_Assertoric_Syllogistic.html --- a/web/entries/Aristotles_Assertoric_Syllogistic.html +++ b/web/entries/Aristotles_Assertoric_Syllogistic.html @@ -1,203 +1,203 @@ Aristotle's Assertoric Syllogistic - Archive of Formal Proofs

 

 

 

 

 

 

Aristotle's Assertoric Syllogistic

 

Title: Aristotle's Assertoric Syllogistic
Author: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki
Submission date: 2019-10-08
Abstract: We formalise with Isabelle/HOL some basic elements of Aristotle's assertoric syllogistic following the article from the Stanford Encyclopedia of Philosophy by Robin Smith. To this end, we use a set theoretic formulation (covering both individual and general predication). In particular, we formalise the deductions in the Figures and after that we present Aristotle's metatheoretical observation that all deductions in the Figures can in fact be reduced to either Barbara or Celarent. As the formal proofs prove to be straightforward, the interest of this entry lies in illustrating the functionality of Isabelle and high efficiency of Sledgehammer for simple exercises in philosophy.
BibTeX:
@article{Aristotles_Assertoric_Syllogistic-AFP,
   author  = {Angeliki Koutsoukou-Argyraki},
   title   = {Aristotle's Assertoric Syllogistic},
   journal = {Archive of Formal Proofs},
   month   = oct,
   year    = 2019,
   note    = {\url{https://isa-afp.org/entries/Aristotles_Assertoric_Syllogistic.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License

\ No newline at end of file diff --git a/web/entries/Complete_Non_Orders.html b/web/entries/Complete_Non_Orders.html --- a/web/entries/Complete_Non_Orders.html +++ b/web/entries/Complete_Non_Orders.html @@ -1,210 +1,210 @@ Complete Non-Orders and Fixed Points - Archive of Formal Proofs

 

 

 

 

 

 

Complete Non-Orders and Fixed Points

 

Title: Complete Non-Orders and Fixed Points
Authors: Akihisa Yamada (akihisa /dot/ yamada /at/ aist /dot/ go /dot/ jp) and Jérémy Dubut
Submission date: 2019-06-27
Abstract: We develop an Isabelle/HOL library of order-theoretic concepts, such as various completeness conditions and fixed-point theorems. We keep our formalization as general as possible: we reprove several well-known results about complete orders, often without any properties of ordering, thus complete non-orders. In particular, we generalize the Knaster–Tarski theorem so that we ensure the existence of a quasi-fixed point of monotone maps over complete non-orders, and show that the set of quasi-fixed points is complete under a mild -condition---attractivity---which is implied by either antisymmetry or +condition—attractivity—which is implied by either antisymmetry or transitivity. This result generalizes and strengthens a result by Stauti and Maaden. Finally, we recover Kleene’s fixed-point theorem for omega-complete non-orders, again using attractivity to prove that Kleene’s fixed points are least quasi-fixed points.
BibTeX:
@article{Complete_Non_Orders-AFP,
   author  = {Akihisa Yamada and Jérémy Dubut},
   title   = {Complete Non-Orders and Fixed Points},
   journal = {Archive of Formal Proofs},
   month   = jun,
   year    = 2019,
   note    = {\url{https://isa-afp.org/entries/Complete_Non_Orders.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License

\ No newline at end of file diff --git a/web/entries/Conditional_Transfer_Rule.html b/web/entries/Conditional_Transfer_Rule.html --- a/web/entries/Conditional_Transfer_Rule.html +++ b/web/entries/Conditional_Transfer_Rule.html @@ -1,187 +1,189 @@ Conditional Transfer Rule - Archive of Formal Proofs

 

 

 

 

 

 

Conditional Transfer Rule

 

- + + +
Title: Conditional Transfer Rule
Author: Mihails Milehins (user9716869 /at/ gmail /dot/ com)
Submission date: 2021-09-06
Abstract: This article provides a collection of experimental utilities for unoverloading of definitions and synthesis of conditional transfer rules for the object logic Isabelle/HOL of the formal proof assistant Isabelle written in Isabelle/ML.
BibTeX:
@article{Conditional_Transfer_Rule-AFP,
   author  = {Mihails Milehins},
   title   = {Conditional Transfer Rule},
   journal = {Archive of Formal Proofs},
   month   = sep,
   year    = 2021,
   note    = {\url{https://isa-afp.org/entries/Conditional_Transfer_Rule.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on:SpecCheck
Used by: Types_To_Sets_Extension

\ No newline at end of file diff --git a/web/entries/Cubic_Quartic_Equations.html b/web/entries/Cubic_Quartic_Equations.html --- a/web/entries/Cubic_Quartic_Equations.html +++ b/web/entries/Cubic_Quartic_Equations.html @@ -1,195 +1,197 @@ Solving Cubic and Quartic Equations - Archive of Formal Proofs

 

 

 

 

 

 

Solving Cubic and Quartic Equations

 

- + + +
Title: Solving Cubic and Quartic Equations
Author: René Thiemann (rene /dot/ thiemann /at/ uibk /dot/ ac /dot/ at)
Submission date: 2021-09-03
Abstract:

We formalize Cardano's formula to solve a cubic equation $$ax^3 + bx^2 + cx + d = 0,$$ as well as Ferrari's formula to solve a quartic equation. We further turn both formulas into executable algorithms based on the algebraic number implementation in the AFP. To this end we also slightly extended this library, namely by making the minimal polynomial of an algebraic number executable, and by defining and implementing $n$-th roots of complex numbers.

BibTeX:
@article{Cubic_Quartic_Equations-AFP,
   author  = {René Thiemann},
   title   = {Solving Cubic and Quartic Equations},
   journal = {Archive of Formal Proofs},
   month   = sep,
   year    = 2021,
   note    = {\url{https://isa-afp.org/entries/Cubic_Quartic_Equations.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on: Algebraic_Numbers, Complex_Geometry
Used by:Factor_Algebraic_Polynomial

\ No newline at end of file diff --git a/web/entries/Factor_Algebraic_Polynomial.html b/web/entries/Factor_Algebraic_Polynomial.html new file mode 100644 --- /dev/null +++ b/web/entries/Factor_Algebraic_Polynomial.html @@ -0,0 +1,200 @@ + + + + +Factorization of Polynomials with Algebraic Coefficients - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + +

 

+

 

+
+
+

 

+

Factorization + + of + + Polynomials + + with + + Algebraic + + Coefficients + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:Factorization of Polynomials with Algebraic Coefficients
+ Authors: + + Manuel Eberl and + René Thiemann (rene /dot/ thiemann /at/ uibk /dot/ ac /dot/ at) +
Submission date:2021-11-08
Abstract: +The AFP already contains a verified implementation of algebraic +numbers. However, it is has a severe limitation in its factorization +algorithm of real and complex polynomials: the factorization is only +guaranteed to succeed if the coefficients of the polynomial are +rational numbers. In this work, we verify an algorithm to factor all +real and complex polynomials whose coefficients are algebraic. The +existence of such an algorithm proves in a constructive way that the +set of complex algebraic numbers is algebraically closed. Internally, +the algorithm is based on resultants of multivariate polynomials and +an approximation algorithm using interval arithmetic.
BibTeX: +
@article{Factor_Algebraic_Polynomial-AFP,
+  author  = {Manuel Eberl and René Thiemann},
+  title   = {Factorization of Polynomials with Algebraic Coefficients},
+  journal = {Archive of Formal Proofs},
+  month   = nov,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/Factor_Algebraic_Polynomial.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
Depends on:Cubic_Quartic_Equations
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/Girth_Chromatic.html b/web/entries/Girth_Chromatic.html --- a/web/entries/Girth_Chromatic.html +++ b/web/entries/Girth_Chromatic.html @@ -1,259 +1,259 @@ A Probabilistic Proof of the Girth-Chromatic Number Theorem - Archive of Formal Proofs

 

 

 

 

 

 

A Probabilistic Proof of the Girth-Chromatic Number Theorem

 

- +
Title: A Probabilistic Proof of the Girth-Chromatic Number Theorem
Author: Lars Noschinski
Submission date: 2012-02-06
Abstract: This works presents a formalization of the Girth-Chromatic number theorem in graph theory, stating that graphs with arbitrarily large girth and chromatic number exist. The proof uses the theory of Random Graphs to prove the existence with probabilistic arguments.
BibTeX:
@article{Girth_Chromatic-AFP,
   author  = {Lars Noschinski},
   title   = {A Probabilistic Proof of the Girth-Chromatic Number Theorem},
   journal = {Archive of Formal Proofs},
   month   = feb,
   year    = 2012,
   note    = {\url{https://isa-afp.org/entries/Girth_Chromatic.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Used by:Random_Graph_Subgraph_Threshold
Random_Graph_Subgraph_Threshold, Szemeredi_Regularity

\ No newline at end of file diff --git a/web/entries/Green.html b/web/entries/Green.html --- a/web/entries/Green.html +++ b/web/entries/Green.html @@ -1,218 +1,218 @@ An Isabelle/HOL formalisation of Green's Theorem - Archive of Formal Proofs

 

 

 

 

 

 

An Isabelle/HOL formalisation of Green's Theorem

 

Title: An Isabelle/HOL formalisation of Green's Theorem
Authors: Mohammad Abdulaziz and Lawrence C. Paulson
Submission date: 2018-01-11
Abstract: -We formalise a statement of Green’s theorem---the first formalisation to -our knowledge---in Isabelle/HOL. The theorem statement that we formalise +We formalise a statement of Green’s theorem—the first formalisation to +our knowledge—in Isabelle/HOL. The theorem statement that we formalise is enough for most applications, especially in physics and engineering. Our formalisation is made possible by a novel proof that avoids the ubiquitous line integral cancellation argument. This eliminates the need to formalise orientations and region boundaries explicitly with respect to the outwards-pointing normal vector. Instead we appeal to a homological argument about equivalences between paths.
BibTeX:
@article{Green-AFP,
   author  = {Mohammad Abdulaziz and Lawrence C. Paulson},
   title   = {An Isabelle/HOL formalisation of Green's Theorem},
   journal = {Archive of Formal Proofs},
   month   = jan,
   year    = 2018,
   note    = {\url{https://isa-afp.org/entries/Green.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License

\ No newline at end of file diff --git a/web/entries/Irrational_Series_Erdos_Straus.html b/web/entries/Irrational_Series_Erdos_Straus.html --- a/web/entries/Irrational_Series_Erdos_Straus.html +++ b/web/entries/Irrational_Series_Erdos_Straus.html @@ -1,212 +1,212 @@ Irrationality Criteria for Series by Erdős and Straus - Archive of Formal Proofs

 

 

 

 

 

 

Irrationality Criteria for Series by Erdős and Straus

 

Title: Irrationality Criteria for Series by Erdős and Straus
Authors: - Angeliki Koutsoukou-Argyraki and + Angeliki Koutsoukou-Argyraki and Wenda Li
Submission date: 2020-05-12
Abstract: We formalise certain irrationality criteria for infinite series of the form: \[\sum_{n=1}^\infty \frac{b_n}{\prod_{i=1}^n a_i} \] where $\{b_n\}$ is a sequence of integers and $\{a_n\}$ a sequence of positive integers with $a_n >1$ for all large n. The results are due to P. Erdős and E. G. Straus [1]. In particular, we formalise Theorem 2.1, Corollary 2.10 and Theorem 3.1. The latter is an application of Theorem 2.1 involving the prime numbers.
BibTeX:
@article{Irrational_Series_Erdos_Straus-AFP,
   author  = {Angeliki Koutsoukou-Argyraki and Wenda Li},
   title   = {Irrationality Criteria for Series by Erdős and Straus},
   journal = {Archive of Formal Proofs},
   month   = may,
   year    = 2020,
   note    = {\url{https://isa-afp.org/entries/Irrational_Series_Erdos_Straus.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on: Prime_Distribution_Elementary, Prime_Number_Theorem

\ No newline at end of file diff --git a/web/entries/Irrationality_J_Hancl.html b/web/entries/Irrationality_J_Hancl.html --- a/web/entries/Irrationality_J_Hancl.html +++ b/web/entries/Irrationality_J_Hancl.html @@ -1,211 +1,211 @@ Irrational Rapidly Convergent Series - Archive of Formal Proofs

 

 

 

 

 

 

Irrational Rapidly Convergent Series

 

Title: Irrational Rapidly Convergent Series
Authors: - Angeliki Koutsoukou-Argyraki and + Angeliki Koutsoukou-Argyraki and Wenda Li
Submission date: 2018-05-23
Abstract: We formalize with Isabelle/HOL a proof of a theorem by J. Hancl asserting the irrationality of the sum of a series consisting of rational numbers, built up by sequences that fulfill certain properties. Even though the criterion is a number theoretic result, the proof makes use only of analytical arguments. We also formalize a corollary of the theorem for a specific series fulfilling the assumptions of the theorem.
BibTeX:
@article{Irrationality_J_Hancl-AFP,
   author  = {Angeliki Koutsoukou-Argyraki and Wenda Li},
   title   = {Irrational Rapidly Convergent Series},
   journal = {Archive of Formal Proofs},
   month   = may,
   year    = 2018,
   note    = {\url{https://isa-afp.org/entries/Irrationality_J_Hancl.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License

\ No newline at end of file diff --git a/web/entries/Octonions.html b/web/entries/Octonions.html --- a/web/entries/Octonions.html +++ b/web/entries/Octonions.html @@ -1,201 +1,201 @@ Octonions - Archive of Formal Proofs

 

 

 

 

 

 

Octonions

 

Title: Octonions
Author: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki
Submission date: 2018-09-14
Abstract: We develop the basic theory of Octonions, including various identities and properties of the octonions and of the octonionic product, a description of 7D isometries and representations of orthogonal transformations. To this end we first develop the theory of the vector cross product in 7 dimensions. The development of the theory of Octonions is inspired by that of the theory of Quaternions by Lawrence Paulson. However, we do not work within the type class real_algebra_1 because the octonionic product is not associative.
BibTeX:
@article{Octonions-AFP,
   author  = {Angeliki Koutsoukou-Argyraki},
   title   = {Octonions},
   journal = {Archive of Formal Proofs},
   month   = sep,
   year    = 2018,
   note    = {\url{https://isa-afp.org/entries/Octonions.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License

\ No newline at end of file diff --git a/web/entries/PAL.html b/web/entries/PAL.html new file mode 100644 --- /dev/null +++ b/web/entries/PAL.html @@ -0,0 +1,202 @@ + + + + +Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + +

 

+

 

+
+
+

 

+

Automating + + Public + + Announcement + + Logic + + and + + the + + Wise + + Men + + Puzzle + + in + + Isabelle/HOL + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL
+ Authors: + + Christoph Benzmüller and + Sebastian Reiche +
Submission date:2021-11-08
Abstract: +We present a shallow embedding of public announcement logic (PAL) with +relativized general knowledge in HOL. We then use PAL to obtain an +elegant encoding of the wise men puzzle, which we solve automatically +using sledgehammer.
BibTeX: +
@article{PAL-AFP,
+  author  = {Christoph Benzmüller and Sebastian Reiche},
+  title   = {Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL},
+  journal = {Archive of Formal Proofs},
+  month   = nov,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/PAL.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/Real_Power.html b/web/entries/Real_Power.html new file mode 100644 --- /dev/null +++ b/web/entries/Real_Power.html @@ -0,0 +1,210 @@ + + + + +Real Exponents as the Limits of Sequences of Rational Exponents - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + +

 

+

 

+
+
+

 

+

Real + + Exponents + + as + + the + + Limits + + of + + Sequences + + of + + Rational + + Exponents + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:Real Exponents as the Limits of Sequences of Rational Exponents
+ Author: + + Jacques D. Fleuriot +
Submission date:2021-11-08
Abstract: +In this formalisation, we construct real exponents as the limits of +sequences of rational exponents. In particular, if $a \ge 1$ and $x +\in \mathbb{R}$, we choose an increasing rational sequence $r_n$ such +that $\lim_{n\to\infty} {r_n} = x$. Then the sequence $a^{r_n}$ is +increasing and if $r$ is any rational number such that $r > x$, +$a^{r_n}$ is bounded above by $a^r$. By the convergence criterion for +monotone sequences, $a^{r_n}$ converges. We define $a^ x = +\lim_{n\to\infty} a^{r_n}$ and show that it has the expected +properties (for $a \ge 0$). This particular construction of real +exponents is needed instead of the usual one using the natural +logarithm and exponential functions (which already exists in Isabelle) +to support our mechanical derivation of Euler's exponential +series as an ``infinite polynomial". Aside from helping us avoid +circular reasoning, this is, as far as we are aware, the first time +real exponents are mechanised in this way within a proof assistant.
BibTeX: +
@article{Real_Power-AFP,
+  author  = {Jacques D. Fleuriot},
+  title   = {Real Exponents as the Limits of Sequences of Rational Exponents},
+  journal = {Archive of Formal Proofs},
+  month   = nov,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/Real_Power.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/Schutz_Spacetime.html b/web/entries/Schutz_Spacetime.html --- a/web/entries/Schutz_Spacetime.html +++ b/web/entries/Schutz_Spacetime.html @@ -1,196 +1,196 @@ Schutz' Independent Axioms for Minkowski Spacetime - Archive of Formal Proofs

 

 

 

 

 

 

Schutz' Independent Axioms for Minkowski Spacetime

 

Title: Schutz' Independent Axioms for Minkowski Spacetime
Authors: Richard Schmoetten (s1311325 /at/ sms /dot/ ed /dot/ ac /dot/ uk), Jake Palmer (jake /dot/ palmer /at/ ed /dot/ ac /dot/ uk) and - Jacques Fleuriot ( jdf /at/ ed /dot/ ac /dot/ uk) + Jacques Fleuriot
Submission date: 2021-07-27
Abstract: This is a formalisation of Schutz' system of axioms for Minkowski spacetime published under the name "Independent axioms for Minkowski space-time" in 1997, as well as most of the results in the third chapter ("Temporal Order on a Path") of the above monograph. Many results are proven here that cannot be found in Schutz, either preceding the theorem they are needed for, or within their own thematic section.
BibTeX:
@article{Schutz_Spacetime-AFP,
   author  = {Richard Schmoetten and Jake Palmer and Jacques Fleuriot},
   title   = {Schutz' Independent Axioms for Minkowski Spacetime},
   journal = {Archive of Formal Proofs},
   month   = jul,
   year    = 2021,
   note    = {\url{https://isa-afp.org/entries/Schutz_Spacetime.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License

\ No newline at end of file diff --git a/web/entries/SpecCheck.html b/web/entries/SpecCheck.html --- a/web/entries/SpecCheck.html +++ b/web/entries/SpecCheck.html @@ -1,201 +1,201 @@ SpecCheck - Specification-Based Testing for Isabelle/ML - Archive of Formal Proofs

 

 

 

 

 

 

SpecCheck - Specification-Based Testing for Isabelle/ML

 

- +
Title: SpecCheck - Specification-Based Testing for Isabelle/ML
Authors: Kevin Kappelmann, Lukas Bulwahn (lukas /dot/ bulwahn /at/ gmail /dot/ com) and Sebastian Willenbrink (sebastian /dot/ willenbrink /at/ tum /dot/ de)
Submission date: 2021-07-01
Abstract: SpecCheck is a QuickCheck-like testing framework for Isabelle/ML. You can use it to write specifications for ML functions. SpecCheck then checks whether your specification holds by testing your function against a given number of generated inputs. It helps you to identify bugs by printing counterexamples on failure and provides you timing information. SpecCheck is customisable and allows you to specify your own input generators, test output formats, as well as pretty printers and shrinking functions for counterexamples among other things.
BibTeX:
@article{SpecCheck-AFP,
   author  = {Kevin Kappelmann and Lukas Bulwahn and Sebastian Willenbrink},
   title   = {SpecCheck - Specification-Based Testing for Isabelle/ML},
   journal = {Archive of Formal Proofs},
   month   = jul,
   year    = 2021,
   note    = {\url{https://isa-afp.org/entries/SpecCheck.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Used by:Regex_Equivalence
Conditional_Transfer_Rule, Regex_Equivalence, Types_To_Sets_Extension

\ No newline at end of file diff --git a/web/entries/Szemeredi_Regularity.html b/web/entries/Szemeredi_Regularity.html new file mode 100644 --- /dev/null +++ b/web/entries/Szemeredi_Regularity.html @@ -0,0 +1,198 @@ + + + + +Szemerédi's Regularity Lemma - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + +

 

+

 

+
+
+

 

+

Szemerédi's + + Regularity + + Lemma + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:Szemerédi's Regularity Lemma
+ Authors: + + Chelsea Edmonds, + Angeliki Koutsoukou-Argyraki and + Lawrence C. Paulson +
Submission date:2021-11-05
Abstract: +Szemerédi's +regularity lemma is a key result in the study of large +graphs. It asserts the existence an upper bound on the number of parts +the vertices of a graph need to be partitioned into such that the +edges between the parts are random in a certain sense. This bound +depends only on the desired precision and not on the graph itself, in +the spirit of Ramsey's theorem. The formalisation follows online +course notes by Tim +Gowers and Yufei +Zhao.
BibTeX: +
@article{Szemeredi_Regularity-AFP,
+  author  = {Chelsea Edmonds and Angeliki Koutsoukou-Argyraki and Lawrence C. Paulson},
+  title   = {Szemerédi's Regularity Lemma},
+  journal = {Archive of Formal Proofs},
+  month   = nov,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/Szemeredi_Regularity.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
Depends on:Girth_Chromatic
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/Transcendence_Series_Hancl_Rucki.html b/web/entries/Transcendence_Series_Hancl_Rucki.html --- a/web/entries/Transcendence_Series_Hancl_Rucki.html +++ b/web/entries/Transcendence_Series_Hancl_Rucki.html @@ -1,212 +1,212 @@ The Transcendence of Certain Infinite Series - Archive of Formal Proofs

 

 

 

 

 

 

The Transcendence of Certain Infinite Series

 

Title: The Transcendence of Certain Infinite Series
Authors: - Angeliki Koutsoukou-Argyraki and + Angeliki Koutsoukou-Argyraki and Wenda Li
Submission date: 2019-03-27
Abstract: We formalize the proofs of two transcendence criteria by J. Hančl and P. Rucki that assert the transcendence of the sums of certain infinite series built up by sequences that fulfil certain properties. Both proofs make use of Roth's celebrated theorem on diophantine approximations to algebraic numbers from 1955 which we implement as an assumption without having formalised its proof.
BibTeX:
@article{Transcendence_Series_Hancl_Rucki-AFP,
   author  = {Angeliki Koutsoukou-Argyraki and Wenda Li},
   title   = {The Transcendence of Certain Infinite Series},
   journal = {Archive of Formal Proofs},
   month   = mar,
   year    = 2019,
   note    = {\url{https://isa-afp.org/entries/Transcendence_Series_Hancl_Rucki.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on: Prime_Number_Theorem

\ No newline at end of file diff --git a/web/entries/Types_To_Sets_Extension.html b/web/entries/Types_To_Sets_Extension.html --- a/web/entries/Types_To_Sets_Extension.html +++ b/web/entries/Types_To_Sets_Extension.html @@ -1,205 +1,205 @@ Extension of Types-To-Sets - Archive of Formal Proofs

 

 

 

 

 

 

Extension of Types-To-Sets

 

- +
Title: Extension of Types-To-Sets
Author: Mihails Milehins (user9716869 /at/ gmail /dot/ com)
Submission date: 2021-09-06
Abstract: In their article titled From Types to Sets by Local Type Definitions in Higher-Order Logic and published in the proceedings of the conference Interactive Theorem Proving in 2016, Ondřej Kunčar and Andrei Popescu propose an extension of the logic Isabelle/HOL and an associated algorithm for the relativization of the type-based theorems to more flexible set-based theorems, collectively referred to as Types-To-Sets. One of the aims of their work was to open an opportunity for the development of a software tool for applied relativization in the implementation of the logic Isabelle/HOL of the proof assistant Isabelle. In this article, we provide a prototype of a software framework for the interactive automated relativization of theorems in Isabelle/HOL, developed as an extension of the proof language Isabelle/Isar. The software framework incorporates the implementation of the proposed extension of the logic, and builds upon some of the ideas for further work expressed in the original article on Types-To-Sets by Ondřej Kunčar and Andrei Popescu and the subsequent article Smooth Manifolds and Types to Sets for Linear Algebra in Isabelle/HOL that was written by Fabian Immler and Bohua Zhan and published in the proceedings of the International Conference on Certified Programs and Proofs in 2019.
BibTeX:
@article{Types_To_Sets_Extension-AFP,
   author  = {Mihails Milehins},
   title   = {Extension of Types-To-Sets},
   journal = {Archive of Formal Proofs},
   month   = sep,
   year    = 2021,
   note    = {\url{https://isa-afp.org/entries/Types_To_Sets_Extension.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on:Conditional_Transfer_Rule
Conditional_Transfer_Rule, SpecCheck

\ No newline at end of file diff --git a/web/index.html b/web/index.html --- a/web/index.html +++ b/web/index.html @@ -1,5810 +1,5846 @@ Archive of Formal Proofs

 

 

 

 

 

 

Archive of Formal Proofs

 

The Archive of Formal Proofs is a collection of proof libraries, examples, and larger scientific developments, mechanically checked in the theorem prover Isabelle. It is organized in the way of a scientific journal, is indexed by dblp and has an ISSN: 2150-914x. Submissions are refereed. The preferred citation style is available [here]. We encourage companion AFP submissions to conference and journal publications.

A development version of the archive is available as well.

 

 

+ + + + + + + + + + + +
2021
+ 2021-11-08: Real Exponents as the Limits of Sequences of Rational Exponents +
+ Author: + Jacques D. Fleuriot +
+ 2021-11-08: Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL +
+ Authors: + Christoph Benzmüller + and Sebastian Reiche +
+ 2021-11-08: Factorization of Polynomials with Algebraic Coefficients +
+ Authors: + Manuel Eberl + and René Thiemann +
+ 2021-11-05: Szemerédi's Regularity Lemma +
+ Authors: + Chelsea Edmonds, + Angeliki Koutsoukou-Argyraki + and Lawrence C. Paulson +
2021-10-28: Quantum and Classical Registers
Author: Dominique Unruh
2021-10-19: Belief Revision Theory
Authors: Valentin Fouillard, Safouan Taha, Frédéric Boulanger and Nicolas Sabouret
2021-10-13: X86 instruction semantics and basic block symbolic execution
Authors: Freek Verbeek, Abhijith Bharadwaj, Joshua Bockenek, Ian Roessle, Timmy Weerwag and Binoy Ravindran
2021-10-12: Algebras for Iteration, Infinite Executions and Correctness of Sequential Computations
Author: Walter Guttmann
2021-10-02: Verified Quadratic Virtual Substitution for Real Arithmetic
Authors: Matias Scharager, Katherine Cordwell, Stefan Mitsch and André Platzer
2021-09-24: Soundness and Completeness of an Axiomatic System for First-Order Logic
Author: Asta Halkjær From
2021-09-18: Complex Bounded Operators
Authors: Jose Manuel Rodriguez Caballero and Dominique Unruh
2021-09-16: A Formalization of Weighted Path Orders and Recursive Path Orders
Authors: Christian Sternagel, René Thiemann and Akihisa Yamada
2021-09-06: Extension of Types-To-Sets
Author: Mihails Milehins
2021-09-06: IDE: Introduction, Destruction, Elimination
Author: Mihails Milehins
2021-09-06: Conditional Transfer Rule
Author: Mihails Milehins
2021-09-06: Conditional Simplification
Author: Mihails Milehins
2021-09-06: Category Theory for ZFC in HOL III: Universal Constructions
Author: Mihails Milehins
2021-09-06: Category Theory for ZFC in HOL I: Foundations: Design Patterns, Set Theory, Digraphs, Semicategories
Author: Mihails Milehins
2021-09-06: Category Theory for ZFC in HOL II: Elementary Theory of 1-Categories
Author: Mihails Milehins
2021-09-05: A data flow analysis algorithm for computing dominators
Author: Nan Jiang
2021-09-03: Solving Cubic and Quartic Equations
Author: René Thiemann
2021-08-26: Logging-independent Message Anonymity in the Relational Method
Author: Pasquale Noce
2021-08-21: The Theorem of Three Circles
Authors: Fox Thomson and Wenda Li
2021-08-16: Fresh identifiers
Authors: Andrei Popescu and Thomas Bauereiss
2021-08-16: CoSMed: A confidentiality-verified social media platform
Authors: Thomas Bauereiss and Andrei Popescu
2021-08-16: CoSMeDis: A confidentiality-verified distributed social media platform
Authors: Thomas Bauereiss and Andrei Popescu
2021-08-16: CoCon: A Confidentiality-Verified Conference Management System
Authors: Andrei Popescu, Peter Lammich and Thomas Bauereiss
2021-08-16: Compositional BD Security
Authors: Thomas Bauereiss and Andrei Popescu
2021-08-13: Combinatorial Design Theory
Authors: Chelsea Edmonds and Lawrence Paulson
2021-08-03: Relational Forests
Author: Walter Guttmann
2021-07-27: Schutz' Independent Axioms for Minkowski Spacetime
Authors: Richard Schmoetten, Jake Palmer - and Jacques Fleuriot + and Jacques Fleuriot
2021-07-07: Finitely Generated Abelian Groups
Authors: Joseph Thommes and Manuel Eberl
2021-07-01: SpecCheck - Specification-Based Testing for Isabelle/ML
Authors: Kevin Kappelmann, Lukas Bulwahn and Sebastian Willenbrink
2021-06-22: Van der Waerden's Theorem
Authors: Katharina Kreuzer and Manuel Eberl
2021-06-18: MiniSail - A kernel language for the ISA specification language SAIL
Author: Mark Wassell
2021-06-17: Public Announcement Logic
Author: Asta Halkjær From
2021-06-04: A Shorter Compiler Correctness Proof for Language IMP
Author: Pasquale Noce
2021-05-24: Lyndon words
Authors: Štěpán Holub and Štěpán Starosta
2021-05-24: Graph Lemma
Authors: Štěpán Holub and Štěpán Starosta
2021-05-24: Combinatorics on Words Basics
Authors: Štěpán Holub, Martin Raška and Štěpán Starosta
2021-04-30: Regression Test Selection
Author: Susannah Mansky
2021-04-27: Isabelle's Metalogic: Formalization and Proof Checker
Authors: Tobias Nipkow and Simon Roßkopf
2021-04-27: Lifting the Exponent
Author: Jakub Kądziołka
2021-04-24: The BKR Decision Procedure for Univariate Real Arithmetic
Authors: Katherine Cordwell, Yong Kiam Tan and André Platzer
2021-04-23: Gale-Stewart Games
Author: Sebastiaan Joosten
2021-04-13: Formalization of Timely Dataflow's Progress Tracking Protocol
Authors: Matthias Brun, Sára Decova, Andrea Lattuada and Dmitriy Traytel
2021-04-01: Information Flow Control via Dependency Tracking
Author: Benedikt Nordhoff
2021-03-29: Grothendieck's Schemes in Algebraic Geometry
Authors: Anthony Bordg, Lawrence Paulson and Wenda Li
2021-03-23: Hensel's Lemma for the p-adic Integers
Author: Aaron Crighton
2021-03-17: Constructive Cryptography in HOL: the Communication Modeling Aspect
Authors: Andreas Lochbihler and S. Reza Sefidgar
2021-03-12: Two algorithms based on modular arithmetic: lattice basis reduction and Hermite normal form computation
Authors: Ralph Bottesch, Jose Divasón and René Thiemann
2021-03-03: Quantum projective measurements and the CHSH inequality
Author: Mnacho Echenim
2021-03-03: The Hermite–Lindemann–Weierstraß Transcendence Theorem
Author: Manuel Eberl
2021-03-01: Mereology
Author: Ben Blumson
2021-02-25: The Sunflower Lemma of Erdős and Rado
Author: René Thiemann
2021-02-24: A Verified Imperative Implementation of B-Trees
Author: Niels Mündler
2021-02-17: Formal Puiseux Series
Author: Manuel Eberl
2021-02-10: The Laws of Large Numbers
Author: Manuel Eberl
2021-01-31: Tarski's Parallel Postulate implies the 5th Postulate of Euclid, the Postulate of Playfair and the original Parallel Postulate of Euclid
Author: Roland Coghetto
2021-01-30: Solution to the xkcd Blue Eyes puzzle
Author: Jakub Kądziołka
2021-01-18: Hood-Melville Queue
Author: Alejandro Gómez-Londoño
2021-01-11: JinjaDCI: a Java semantics with dynamic class initialization
Author: Susannah Mansky

 

2020
2020-12-27: Cofinality and the Delta System Lemma
Author: Pedro Sánchez Terraf
2020-12-17: Topological semantics for paraconsistent and paracomplete logics
Author: David Fuenmayor
2020-12-08: Relational Minimum Spanning Tree Algorithms
Authors: Walter Guttmann and Nicolas Robinson-O'Brien
2020-12-07: Inline Caching and Unboxing Optimization for Interpreters
Author: Martin Desharnais
2020-12-05: The Relational Method with Message Anonymity for the Verification of Cryptographic Protocols
Author: Pasquale Noce
2020-11-22: Isabelle Marries Dirac: a Library for Quantum Computation and Quantum Information
Authors: Anthony Bordg, Hanna Lachnitt and Yijun He
2020-11-19: The HOL-CSP Refinement Toolkit
Authors: Safouan Taha, Burkhart Wolff and Lina Ye
2020-10-29: Verified SAT-Based AI Planning
Authors: Mohammad Abdulaziz and Friedrich Kurz
2020-10-29: AI Planning Languages Semantics
Authors: Mohammad Abdulaziz and Peter Lammich
2020-10-20: A Sound Type System for Physical Quantities, Units, and Measurements
Authors: Simon Foster and Burkhart Wolff
2020-10-12: Finite Map Extras
Author: Javier Díaz
2020-09-28: A Formal Model of the Safely Composable Document Object Model with Shadow Roots
Authors: Achim D. Brucker and Michael Herzberg
2020-09-28: A Formal Model of the Document Object Model with Shadow Roots
Authors: Achim D. Brucker and Michael Herzberg
2020-09-28: A Formalization of Safely Composable Web Components
Authors: Achim D. Brucker and Michael Herzberg
2020-09-28: A Formalization of Web Components
Authors: Achim D. Brucker and Michael Herzberg
2020-09-28: The Safely Composable DOM
Authors: Achim D. Brucker and Michael Herzberg
2020-09-16: Syntax-Independent Logic Infrastructure
Authors: Andrei Popescu and Dmitriy Traytel
2020-09-16: Robinson Arithmetic
Authors: Andrei Popescu and Dmitriy Traytel
2020-09-16: An Abstract Formalization of Gödel's Incompleteness Theorems
Authors: Andrei Popescu and Dmitriy Traytel
2020-09-16: From Abstract to Concrete Gödel's Incompleteness Theorems—Part II
Authors: Andrei Popescu and Dmitriy Traytel
2020-09-16: From Abstract to Concrete Gödel's Incompleteness Theorems—Part I
Authors: Andrei Popescu and Dmitriy Traytel
2020-09-07: A Formal Model of Extended Finite State Machines
Authors: Michael Foster, Achim D. Brucker, Ramsay G. Taylor and John Derrick
2020-09-07: Inference of Extended Finite State Machines
Authors: Michael Foster, Achim D. Brucker, Ramsay G. Taylor and John Derrick
2020-08-31: Practical Algebraic Calculus Checker
Authors: Mathias Fleury and Daniela Kaufmann
2020-08-31: Some classical results in inductive inference of recursive functions
Author: Frank J. Balbach
2020-08-26: Relational Disjoint-Set Forests
Author: Walter Guttmann
2020-08-25: Extensions to the Comprehensive Framework for Saturation Theorem Proving
Authors: Jasmin Blanchette and Sophie Tourret
2020-08-25: Putting the `K' into Bird's derivation of Knuth-Morris-Pratt string matching
Author: Peter Gammie
2020-08-04: Amicable Numbers
Author: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki
2020-08-03: Ordinal Partitions
Author: Lawrence C. Paulson
2020-07-21: A Formal Proof of The Chandy--Lamport Distributed Snapshot Algorithm
Authors: Ben Fiedler and Dmitriy Traytel
2020-07-13: Relational Characterisations of Paths
Authors: Walter Guttmann and Peter Höfner
2020-06-01: A Formally Verified Checker of the Safe Distance Traffic Rules for Autonomous Vehicles
Authors: Albert Rizaldi and Fabian Immler
2020-05-23: A verified algorithm for computing the Smith normal form of a matrix
Author: Jose Divasón
2020-05-16: The Nash-Williams Partition Theorem
Author: Lawrence C. Paulson
2020-05-13: A Formalization of Knuth–Bendix Orders
Authors: Christian Sternagel and René Thiemann
2020-05-12: Irrationality Criteria for Series by Erdős and Straus
Authors: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki and Wenda Li
2020-05-11: Recursion Theorem in ZF
Author: Georgy Dunaev
2020-05-08: An Efficient Normalisation Procedure for Linear Temporal Logic: Isabelle/HOL Formalisation
Author: Salomon Sickert
2020-05-06: Formalization of Forcing in Isabelle/ZF
Authors: Emmanuel Gunther, Miguel Pagano and Pedro Sánchez Terraf
2020-05-02: Banach-Steinhaus Theorem
Authors: Dominique Unruh and Jose Manuel Rodriguez Caballero
2020-04-27: Attack Trees in Isabelle for GDPR compliance of IoT healthcare systems
Author: Florian Kammueller
2020-04-24: Power Sum Polynomials
Author: Manuel Eberl
2020-04-24: The Lambert W Function on the Reals
Author: Manuel Eberl
2020-04-24: Gaussian Integers
Author: Manuel Eberl
2020-04-19: Matrices for ODEs
Author: Jonathan Julian Huerta y Munive
2020-04-16: Authenticated Data Structures As Functors
Authors: Andreas Lochbihler and Ognjen Marić
2020-04-10: Formalization of an Algorithm for Greedily Computing Associative Aggregations on Sliding Windows
Authors: Lukas Heimes, Dmitriy Traytel and Joshua Schneider
2020-04-09: A Comprehensive Framework for Saturation Theorem Proving
Author: Sophie Tourret
2020-04-09: Formalization of an Optimized Monitoring Algorithm for Metric First-Order Dynamic Logic with Aggregations
Authors: Thibault Dardinier, Lukas Heimes, Martin Raszyk, Joshua Schneider and Dmitriy Traytel
2020-04-08: Stateful Protocol Composition and Typing
Authors: Andreas V. Hess, Sebastian Mödersheim and Achim D. Brucker
2020-04-08: Automated Stateful Protocol Verification
Authors: Andreas V. Hess, Sebastian Mödersheim, Achim D. Brucker and Anders Schlichtkrull
2020-04-07: Lucas's Theorem
Author: Chelsea Edmonds
2020-03-25: Strong Eventual Consistency of the Collaborative Editing Framework WOOT
Authors: Emin Karayel and Edgar Gonzàlez
2020-03-22: Furstenberg's topology and his proof of the infinitude of primes
Author: Manuel Eberl
2020-03-12: An Under-Approximate Relational Logic
Author: Toby Murray
2020-03-07: Hello World
Authors: Cornelius Diekmann and Lars Hupel
2020-02-21: Implementing the Goodstein Function in λ-Calculus
Author: Bertram Felgenhauer
2020-02-10: A Generic Framework for Verified Compilers
Author: Martin Desharnais
2020-02-01: Arithmetic progressions and relative primes
Author: José Manuel Rodríguez Caballero
2020-01-31: A Hierarchy of Algebras for Boolean Subsets
Authors: Walter Guttmann and Bernhard Möller
2020-01-17: Mersenne primes and the Lucas–Lehmer test
Author: Manuel Eberl
2020-01-16: Verified Approximation Algorithms
Authors: Robin Eßmann, Tobias Nipkow, Simon Robillard and Ujkan Sulejmani
2020-01-13: Closest Pair of Points Algorithms
Authors: Martin Rau and Tobias Nipkow
2020-01-09: Skip Lists
Authors: Max W. Haslbeck and Manuel Eberl
2020-01-06: Bicategories
Author: Eugene W. Stark

 

2019
2019-12-27: The Irrationality of ζ(3)
Author: Manuel Eberl
2019-12-20: Formalizing a Seligman-Style Tableau System for Hybrid Logic
Author: Asta Halkjær From
2019-12-18: The Poincaré-Bendixson Theorem
Authors: Fabian Immler and Yong Kiam Tan
2019-12-16: Poincaré Disc Model
Authors: Danijela Simić, Filip Marić and Pierre Boutry
2019-12-16: Complex Geometry
Authors: Filip Marić and Danijela Simić
2019-12-10: Gauss Sums and the Pólya–Vinogradov Inequality
Authors: Rodrigo Raya and Manuel Eberl
2019-12-04: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges
Author: Pasquale Noce
2019-11-27: Interval Arithmetic on 32-bit Words
Author: Brandon Bohrer
2019-10-24: Zermelo Fraenkel Set Theory in Higher-Order Logic
Author: Lawrence C. Paulson
2019-10-22: Isabelle/C
Authors: Frédéric Tuong and Burkhart Wolff
2019-10-16: VerifyThis 2019 -- Polished Isabelle Solutions
Authors: Peter Lammich and Simon Wimmer
2019-10-08: Aristotle's Assertoric Syllogistic
Author: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki
2019-10-07: Sigma Protocols and Commitment Schemes
Authors: David Butler and Andreas Lochbihler
2019-10-04: Clean - An Abstract Imperative Programming Language and its Theory
Authors: Frédéric Tuong and Burkhart Wolff
2019-09-16: Formalization of Multiway-Join Algorithms
Author: Thibault Dardinier
2019-09-10: Verification Components for Hybrid Systems
Author: Jonathan Julian Huerta y Munive
2019-09-06: Fourier Series
Author: Lawrence C Paulson
2019-08-30: A Case Study in Basic Algebra
Author: Clemens Ballarin
2019-08-16: Formalisation of an Adaptive State Counting Algorithm
Author: Robert Sachtleben
2019-08-14: Laplace Transform
Author: Fabian Immler
2019-08-06: Linear Programming
Authors: Julian Parsert and Cezary Kaliszyk
2019-08-06: Communicating Concurrent Kleene Algebra for Distributed Systems Specification
Authors: Maxime Buyse and Jason Jaskolka
2019-08-05: Selected Problems from the International Mathematical Olympiad 2019
Author: Manuel Eberl
2019-08-01: Stellar Quorum Systems
Author: Giuliano Losa
2019-07-30: A Formal Development of a Polychronous Polytimed Coordination Language
Authors: Hai Nguyen Van, Frédéric Boulanger and Burkhart Wolff
2019-07-27: Order Extension and Szpilrajn's Extension Theorem
Authors: Peter Zeller and Lukas Stevens
2019-07-18: A Sequent Calculus for First-Order Logic
Author: Asta Halkjær From
2019-07-08: A Verified Code Generator from Isabelle/HOL to CakeML
Author: Lars Hupel
2019-07-04: Formalization of a Monitoring Algorithm for Metric First-Order Temporal Logic
Authors: Joshua Schneider and Dmitriy Traytel
2019-06-27: Complete Non-Orders and Fixed Points
Authors: Akihisa Yamada and Jérémy Dubut
2019-06-25: Priority Search Trees
Authors: Peter Lammich and Tobias Nipkow
2019-06-25: Purely Functional, Simple, and Efficient Implementation of Prim and Dijkstra
Authors: Peter Lammich and Tobias Nipkow
2019-06-21: Linear Inequalities
Authors: Ralph Bottesch, Alban Reynaud and René Thiemann
2019-06-16: Hilbert's Nullstellensatz
Author: Alexander Maletzky
2019-06-15: Gröbner Bases, Macaulay Matrices and Dubé's Degree Bounds
Author: Alexander Maletzky
2019-06-13: Binary Heaps for IMP2
Author: Simon Griebel
2019-06-03: Differential Game Logic
Author: André Platzer
2019-05-30: Multidimensional Binary Search Trees
Author: Martin Rau
2019-05-14: Formalization of Generic Authenticated Data Structures
Authors: Matthias Brun and Dmitriy Traytel
2019-05-09: Multi-Party Computation
Authors: David Aspinall and David Butler
2019-04-26: HOL-CSP Version 2.0
Authors: Safouan Taha, Lina Ye and Burkhart Wolff
2019-04-16: A Compositional and Unified Translation of LTL into ω-Automata
Authors: Benedikt Seidl and Salomon Sickert
2019-04-06: A General Theory of Syntax with Bindings
Authors: Lorenzo Gheri and Andrei Popescu
2019-03-27: The Transcendence of Certain Infinite Series
Authors: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki and Wenda Li
2019-03-24: Quantum Hoare Logic
Authors: Junyi Liu, Bohua Zhan, Shuling Wang, Shenggang Ying, Tao Liu, Yangjia Li, Mingsheng Ying and Naijun Zhan
2019-03-09: Safe OCL
Author: Denis Nikiforov
2019-02-21: Elementary Facts About the Distribution of Primes
Author: Manuel Eberl
2019-02-14: Kruskal's Algorithm for Minimum Spanning Forest
Authors: Maximilian P.L. Haslbeck, Peter Lammich and Julian Biendarra
2019-02-11: Probabilistic Primality Testing
Authors: Daniel Stüwe and Manuel Eberl
2019-02-08: Universal Turing Machine
Authors: Jian Xu, Xingyuan Zhang, Christian Urban and Sebastiaan J. C. Joosten
2019-02-01: Isabelle/UTP: Mechanised Theory Engineering for Unifying Theories of Programming
Authors: Simon Foster, Frank Zeyda, Yakoub Nemouchi, Pedro Ribeiro and Burkhart Wolff
2019-02-01: The Inversions of a List
Author: Manuel Eberl
2019-01-17: Farkas' Lemma and Motzkin's Transposition Theorem
Authors: Ralph Bottesch, Max W. Haslbeck and René Thiemann
2019-01-15: IMP2 – Simple Program Verification in Isabelle/HOL
Authors: Peter Lammich and Simon Wimmer
2019-01-15: An Algebra for Higher-Order Terms
Author: Lars Hupel
2019-01-07: A Reduction Theorem for Store Buffers
Authors: Ernie Cohen and Norbert Schirmer

 

2018
2018-12-26: A Formal Model of the Document Object Model
Authors: Achim D. Brucker and Michael Herzberg
2018-12-25: Formalization of Concurrent Revisions
Author: Roy Overbeek
2018-12-21: Verifying Imperative Programs using Auto2
Author: Bohua Zhan
2018-12-17: Constructive Cryptography in HOL
Authors: Andreas Lochbihler and S. Reza Sefidgar
2018-12-11: Transformer Semantics
Author: Georg Struth
2018-12-11: Quantales
Author: Georg Struth
2018-12-11: Properties of Orderings and Lattices
Author: Georg Struth
2018-11-23: Graph Saturation
Author: Sebastiaan J. C. Joosten
2018-11-23: A Verified Functional Implementation of Bachmair and Ganzinger's Ordered Resolution Prover
Authors: Anders Schlichtkrull, Jasmin Christian Blanchette and Dmitriy Traytel
2018-11-20: Auto2 Prover
Author: Bohua Zhan
2018-11-16: Matroids
Author: Jonas Keinholz
2018-11-06: Deriving generic class instances for datatypes
Authors: Jonas Rädle and Lars Hupel
2018-10-30: Formalisation and Evaluation of Alan Gewirth's Proof for the Principle of Generic Consistency in Isabelle/HOL
Authors: David Fuenmayor and Christoph Benzmüller
2018-10-29: Epistemic Logic: Completeness of Modal Logics
Author: Asta Halkjær From
2018-10-22: Smooth Manifolds
Authors: Fabian Immler and Bohua Zhan
2018-10-19: Randomised Binary Search Trees
Author: Manuel Eberl
2018-10-19: Formalization of the Embedding Path Order for Lambda-Free Higher-Order Terms
Author: Alexander Bentkamp
2018-10-12: Upper Bounding Diameters of State Spaces of Factored Transition Systems
Authors: Friedrich Kurz and Mohammad Abdulaziz
2018-09-28: The Transcendence of π
Author: Manuel Eberl
2018-09-25: Symmetric Polynomials
Author: Manuel Eberl
2018-09-20: Signature-Based Gröbner Basis Algorithms
Author: Alexander Maletzky
2018-09-19: The Prime Number Theorem
Authors: Manuel Eberl and Lawrence C. Paulson
2018-09-15: Aggregation Algebras
Author: Walter Guttmann
2018-09-14: Octonions
Author: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki
2018-09-05: Quaternions
Author: Lawrence C. Paulson
2018-09-02: The Budan-Fourier Theorem and Counting Real Roots with Multiplicity
Author: Wenda Li
2018-08-24: An Incremental Simplex Algorithm with Unsatisfiable Core Generation
Authors: Filip Marić, Mirko Spasić and René Thiemann
2018-08-14: Minsky Machines
Author: Bertram Felgenhauer
2018-07-16: Pricing in discrete financial models
Author: Mnacho Echenim
2018-07-04: Von-Neumann-Morgenstern Utility Theorem
Authors: Julian Parsert and Cezary Kaliszyk
2018-06-23: Pell's Equation
Author: Manuel Eberl
2018-06-14: Projective Geometry
Author: Anthony Bordg
2018-06-14: The Localization of a Commutative Ring
Author: Anthony Bordg
2018-06-05: Partial Order Reduction
Author: Julian Brunner
2018-05-27: Optimal Binary Search Trees
Authors: Tobias Nipkow and Dániel Somogyi
2018-05-25: Hidden Markov Models
Author: Simon Wimmer
2018-05-24: Probabilistic Timed Automata
Authors: Simon Wimmer and Johannes Hölzl
2018-05-23: Irrational Rapidly Convergent Series
Authors: - Angeliki Koutsoukou-Argyraki + Angeliki Koutsoukou-Argyraki and Wenda Li
2018-05-23: Axiom Systems for Category Theory in Free Logic
Authors: Christoph Benzmüller and Dana Scott
2018-05-22: Monadification, Memoization and Dynamic Programming
Authors: Simon Wimmer, Shuwei Hu and Tobias Nipkow
2018-05-10: OpSets: Sequential Specifications for Replicated Datatypes
Authors: Martin Kleppmann, Victor B. F. Gomes, Dominic P. Mulligan and Alastair R. Beresford
2018-05-07: An Isabelle/HOL Formalization of the Modular Assembly Kit for Security Properties
Authors: Oliver Bračevac, Richard Gay, Sylvia Grewe, Heiko Mantel, Henning Sudbrock and Markus Tasch
2018-04-29: WebAssembly
Author: Conrad Watt
2018-04-27: VerifyThis 2018 - Polished Isabelle Solutions
Authors: Peter Lammich and Simon Wimmer
2018-04-24: Bounded Natural Functors with Covariance and Contravariance
Authors: Andreas Lochbihler and Joshua Schneider
2018-03-22: The Incompatibility of Fishburn-Strategyproofness and Pareto-Efficiency
Authors: Felix Brandt, Manuel Eberl, Christian Saile and Christian Stricker
2018-03-13: Weight-Balanced Trees
Authors: Tobias Nipkow and Stefan Dirix
2018-03-12: CakeML
Authors: Lars Hupel and Yu Zhang
2018-03-01: A Theory of Architectural Design Patterns
Author: Diego Marmsoler
2018-02-26: Hoare Logics for Time Bounds
Authors: Maximilian P. L. Haslbeck and Tobias Nipkow
2018-02-06: Treaps
Authors: Maximilian Haslbeck, Manuel Eberl and Tobias Nipkow
2018-02-06: A verified factorization algorithm for integer polynomials with polynomial complexity
Authors: Jose Divasón, Sebastiaan Joosten, René Thiemann and Akihisa Yamada
2018-02-06: First-Order Terms
Authors: Christian Sternagel and René Thiemann
2018-02-06: The Error Function
Author: Manuel Eberl
2018-02-02: A verified LLL algorithm
Authors: Ralph Bottesch, Jose Divasón, Maximilian Haslbeck, Sebastiaan Joosten, René Thiemann and Akihisa Yamada
2018-01-18: Formalization of Bachmair and Ganzinger's Ordered Resolution Prover
Authors: Anders Schlichtkrull, Jasmin Christian Blanchette, Dmitriy Traytel and Uwe Waldmann
2018-01-16: Gromov Hyperbolicity
Author: Sebastien Gouezel
2018-01-11: An Isabelle/HOL formalisation of Green's Theorem
Authors: Mohammad Abdulaziz and Lawrence C. Paulson
2018-01-08: Taylor Models
Authors: Christoph Traut and Fabian Immler

 

2017
2017-12-22: The Falling Factorial of a Sum
Author: Lukas Bulwahn
2017-12-21: The Median-of-Medians Selection Algorithm
Author: Manuel Eberl
2017-12-21: The Mason–Stothers Theorem
Author: Manuel Eberl
2017-12-21: Dirichlet L-Functions and Dirichlet's Theorem
Author: Manuel Eberl
2017-12-19: Operations on Bounded Natural Functors
Authors: Jasmin Christian Blanchette, Andrei Popescu and Dmitriy Traytel
2017-12-18: The string search algorithm by Knuth, Morris and Pratt
Authors: Fabian Hellauer and Peter Lammich
2017-11-22: Stochastic Matrices and the Perron-Frobenius Theorem
Author: René Thiemann
2017-11-09: The IMAP CmRDT
Authors: Tim Jungnickel, Lennart Oldenburg and Matthias Loibl
2017-11-06: Hybrid Multi-Lane Spatial Logic
Author: Sven Linker
2017-10-26: The Kuratowski Closure-Complement Theorem
Authors: Peter Gammie and Gianpaolo Gioiosa
2017-10-19: Transition Systems and Automata
Author: Julian Brunner
2017-10-19: Büchi Complementation
Author: Julian Brunner
2017-10-17: Evaluate Winding Numbers through Cauchy Indices
Author: Wenda Li
2017-10-17: Count the Number of Complex Roots
Author: Wenda Li
2017-10-14: Homogeneous Linear Diophantine Equations
Authors: Florian Messner, Julian Parsert, Jonas Schöpf and Christian Sternagel
2017-10-12: The Hurwitz and Riemann ζ Functions
Author: Manuel Eberl
2017-10-12: Linear Recurrences
Author: Manuel Eberl
2017-10-12: Dirichlet Series
Author: Manuel Eberl
2017-09-21: Computer-assisted Reconstruction and Assessment of E. J. Lowe's Modal Ontological Argument
Authors: David Fuenmayor and Christoph Benzmüller
2017-09-17: Representation and Partial Automation of the Principia Logico-Metaphysica in Isabelle/HOL
Author: Daniel Kirchner
2017-09-06: Anselm's God in Isabelle/HOL
Author: Ben Blumson
2017-09-01: Microeconomics and the First Welfare Theorem
Authors: Julian Parsert and Cezary Kaliszyk
2017-08-20: Root-Balanced Tree
Author: Tobias Nipkow
2017-08-20: Orbit-Stabiliser Theorem with Application to Rotational Symmetries
Author: Jonas Rädle
2017-08-16: The LambdaMu-calculus
Authors: Cristina Matache, Victor B. F. Gomes and Dominic P. Mulligan
2017-07-31: Stewart's Theorem and Apollonius' Theorem
Author: Lukas Bulwahn
2017-07-28: Dynamic Architectures
Author: Diego Marmsoler
2017-07-21: Declarative Semantics for Functional Languages
Author: Jeremy Siek
2017-07-15: HOLCF-Prelude
Authors: Joachim Breitner, Brian Huffman, Neil Mitchell and Christian Sternagel
2017-07-13: Minkowski's Theorem
Author: Manuel Eberl
2017-07-09: Verified Metatheory and Type Inference for a Name-Carrying Simply-Typed Lambda Calculus
Author: Michael Rawson
2017-07-07: A framework for establishing Strong Eventual Consistency for Conflict-free Replicated Datatypes
Authors: Victor B. F. Gomes, Martin Kleppmann, Dominic P. Mulligan and Alastair R. Beresford
2017-07-06: Stone-Kleene Relation Algebras
Author: Walter Guttmann
2017-06-21: Propositional Proof Systems
Authors: Julius Michaelis and Tobias Nipkow
2017-06-13: Partial Semigroups and Convolution Algebras
Authors: Brijesh Dongol, Victor B. F. Gomes, Ian J. Hayes and Georg Struth
2017-06-06: Buffon's Needle Problem
Author: Manuel Eberl
2017-06-01: Formalizing Push-Relabel Algorithms
Authors: Peter Lammich and S. Reza Sefidgar
2017-06-01: Flow Networks and the Min-Cut-Max-Flow Theorem
Authors: Peter Lammich and S. Reza Sefidgar
2017-05-25: Optics
Authors: Simon Foster and Frank Zeyda
2017-05-24: Developing Security Protocols by Refinement
Authors: Christoph Sprenger and Ivano Somaini
2017-05-24: Dictionary Construction
Author: Lars Hupel
2017-05-08: The Floyd-Warshall Algorithm for Shortest Paths
Authors: Simon Wimmer and Peter Lammich
2017-05-05: Probabilistic while loop
Author: Andreas Lochbihler
2017-05-05: Effect polymorphism in higher-order logic
Author: Andreas Lochbihler
2017-05-05: Monad normalisation
Authors: Joshua Schneider, Manuel Eberl and Andreas Lochbihler
2017-05-05: Game-based cryptography in HOL
Authors: Andreas Lochbihler, S. Reza Sefidgar and Bhargav Bhatt
2017-05-05: CryptHOL
Author: Andreas Lochbihler
2017-05-04: Monoidal Categories
Author: Eugene W. Stark
2017-05-01: Types, Tableaus and Gödel’s God in Isabelle/HOL
Authors: David Fuenmayor and Christoph Benzmüller
2017-04-28: Local Lexing
Author: Steven Obua
2017-04-19: Constructor Functions
Author: Lars Hupel
2017-04-18: Lazifying case constants
Author: Lars Hupel
2017-04-06: Subresultants
Authors: Sebastiaan Joosten, René Thiemann and Akihisa Yamada
2017-04-04: Expected Shape of Random Binary Search Trees
Author: Manuel Eberl
2017-03-15: The number of comparisons in QuickSort
Author: Manuel Eberl
2017-03-15: Lower bound on comparison-based sorting algorithms
Author: Manuel Eberl
2017-03-10: The Euler–MacLaurin Formula
Author: Manuel Eberl
2017-02-28: The Group Law for Elliptic Curves
Author: Stefan Berghofer
2017-02-26: Menger's Theorem
Author: Christoph Dittmann
2017-02-13: Differential Dynamic Logic
Author: Brandon Bohrer
2017-02-10: Abstract Soundness
Authors: Jasmin Christian Blanchette, Andrei Popescu and Dmitriy Traytel
2017-02-07: Stone Relation Algebras
Author: Walter Guttmann
2017-01-31: Refining Authenticated Key Agreement with Strong Adversaries
Authors: Joseph Lallemand and Christoph Sprenger
2017-01-24: Bernoulli Numbers
Authors: Lukas Bulwahn and Manuel Eberl
2017-01-17: Minimal Static Single Assignment Form
Authors: Max Wagner and Denis Lohner
2017-01-17: Bertrand's postulate
Authors: Julian Biendarra and Manuel Eberl
2017-01-12: The Transcendence of e
Author: Manuel Eberl
2017-01-08: Formal Network Models and Their Application to Firewall Policies
Authors: Achim D. Brucker, Lukas Brügger and Burkhart Wolff
2017-01-03: Verification of a Diffie-Hellman Password-based Authentication Protocol by Extending the Inductive Method
Author: Pasquale Noce
2017-01-01: First-Order Logic According to Harrison
Authors: Alexander Birch Jensen, Anders Schlichtkrull and Jørgen Villadsen

 

2016
2016-12-30: Concurrent Refinement Algebra and Rely Quotients
Authors: Julian Fell, Ian J. Hayes and Andrius Velykis
2016-12-29: The Twelvefold Way
Author: Lukas Bulwahn
2016-12-20: Proof Strategy Language
Author: Yutaka Nagashima
2016-12-07: Paraconsistency
Authors: Anders Schlichtkrull and Jørgen Villadsen
2016-11-29: COMPLX: A Verification Framework for Concurrent Imperative Programs
Authors: Sidney Amani, June Andronick, Maksym Bortin, Corey Lewis, Christine Rizkallah and Joseph Tuong
2016-11-23: Abstract Interpretation of Annotated Commands
Author: Tobias Nipkow
2016-11-16: Separata: Isabelle tactics for Separation Algebra
Authors: Zhe Hou, David Sanan, Alwen Tiu, Rajeev Gore and Ranald Clouston
2016-11-12: Formalization of Nested Multisets, Hereditary Multisets, and Syntactic Ordinals
Authors: Jasmin Christian Blanchette, Mathias Fleury and Dmitriy Traytel
2016-11-12: Formalization of Knuth–Bendix Orders for Lambda-Free Higher-Order Terms
Authors: Heiko Becker, Jasmin Christian Blanchette, Uwe Waldmann and Daniel Wand
2016-11-10: Expressiveness of Deep Learning
Author: Alexander Bentkamp
2016-10-25: Modal Logics for Nominal Transition Systems
Authors: Tjark Weber, Lars-Henrik Eriksson, Joachim Parrow, Johannes Borgström and Ramunas Gutkovas
2016-10-24: Stable Matching
Author: Peter Gammie
2016-10-21: LOFT — Verified Migration of Linux Firewalls to SDN
Authors: Julius Michaelis and Cornelius Diekmann
2016-10-19: Source Coding Theorem
Authors: Quentin Hibon and Lawrence C. Paulson
2016-10-19: A formal model for the SPARCv8 ISA and a proof of non-interference for the LEON3 processor
Authors: Zhe Hou, David Sanan, Alwen Tiu and Yang Liu
2016-10-14: The Factorization Algorithm of Berlekamp and Zassenhaus
Authors: Jose Divasón, Sebastiaan Joosten, René Thiemann and Akihisa Yamada
2016-10-11: Intersecting Chords Theorem
Author: Lukas Bulwahn
2016-10-05: Lp spaces
Author: Sebastien Gouezel
2016-09-30: Fisher–Yates shuffle
Author: Manuel Eberl
2016-09-29: Allen's Interval Calculus
Author: Fadoua Ghourabi
2016-09-23: Formalization of Recursive Path Orders for Lambda-Free Higher-Order Terms
Authors: Jasmin Christian Blanchette, Uwe Waldmann and Daniel Wand
2016-09-09: Iptables Semantics
Authors: Cornelius Diekmann and Lars Hupel
2016-09-06: A Variant of the Superposition Calculus
Author: Nicolas Peltier
2016-09-06: Stone Algebras
Author: Walter Guttmann
2016-09-01: Stirling's formula
Author: Manuel Eberl
2016-08-31: Routing
Authors: Julius Michaelis and Cornelius Diekmann
2016-08-24: Simple Firewall
Authors: Cornelius Diekmann, Julius Michaelis and Maximilian Haslbeck
2016-08-18: Infeasible Paths Elimination by Symbolic Execution Techniques: Proof of Correctness and Preservation of Paths
Authors: Romain Aissat, Frederic Voisin and Burkhart Wolff
2016-08-12: Formalizing the Edmonds-Karp Algorithm
Authors: Peter Lammich and S. Reza Sefidgar
2016-08-08: The Imperative Refinement Framework
Author: Peter Lammich
2016-08-07: Ptolemy's Theorem
Author: Lukas Bulwahn
2016-07-17: Surprise Paradox
Author: Joachim Breitner
2016-07-14: Pairing Heap
Authors: Hauke Brinkop and Tobias Nipkow
2016-07-05: A Framework for Verifying Depth-First Search Algorithms
Authors: Peter Lammich and René Neumann
2016-07-01: Chamber Complexes, Coxeter Systems, and Buildings
Author: Jeremy Sylvestre
2016-06-30: The Z Property
Authors: Bertram Felgenhauer, Julian Nagele, Vincent van Oostrom and Christian Sternagel
2016-06-30: The Resolution Calculus for First-Order Logic
Author: Anders Schlichtkrull
2016-06-28: IP Addresses
Authors: Cornelius Diekmann, Julius Michaelis and Lars Hupel
2016-06-28: Compositional Security-Preserving Refinement for Concurrent Imperative Programs
Authors: Toby Murray, Robert Sison, Edward Pierzchalski and Christine Rizkallah
2016-06-26: Category Theory with Adjunctions and Limits
Author: Eugene W. Stark
2016-06-26: Cardinality of Multisets
Author: Lukas Bulwahn
2016-06-25: A Dependent Security Type System for Concurrent Imperative Programs
Authors: Toby Murray, Robert Sison, Edward Pierzchalski and Christine Rizkallah
2016-06-21: Catalan Numbers
Author: Manuel Eberl
2016-06-18: Program Construction and Verification Components Based on Kleene Algebra
Authors: Victor B. F. Gomes and Georg Struth
2016-06-13: Conservation of CSP Noninterference Security under Concurrent Composition
Author: Pasquale Noce
2016-06-09: Finite Machine Word Library
Authors: Joel Beeren, Matthew Fernandez, Xin Gao, Gerwin Klein, Rafal Kolanski, Japheth Lim, Corey Lewis, Daniel Matichuk and Thomas Sewell
2016-05-31: Tree Decomposition
Author: Christoph Dittmann
2016-05-24: POSIX Lexing with Derivatives of Regular Expressions
Authors: Fahad Ausaf, Roy Dyckhoff and Christian Urban
2016-05-24: Cardinality of Equivalence Relations
Author: Lukas Bulwahn
2016-05-20: Perron-Frobenius Theorem for Spectral Radius Analysis
Authors: Jose Divasón, Ondřej Kunčar, René Thiemann and Akihisa Yamada
2016-05-20: The meta theory of the Incredible Proof Machine
Authors: Joachim Breitner and Denis Lohner
2016-05-18: A Constructive Proof for FLP
Authors: Benjamin Bisping, Paul-David Brodmann, Tim Jungnickel, Christina Rickmann, Henning Seidler, Anke Stüber, Arno Wilhelm-Weidner, Kirstin Peters and Uwe Nestmann
2016-05-09: A Formal Proof of the Max-Flow Min-Cut Theorem for Countable Networks
Author: Andreas Lochbihler
2016-05-05: Randomised Social Choice Theory
Author: Manuel Eberl
2016-05-04: The Incompatibility of SD-Efficiency and SD-Strategy-Proofness
Author: Manuel Eberl
2016-05-04: Spivey's Generalized Recurrence for Bell Numbers
Author: Lukas Bulwahn
2016-05-02: Gröbner Bases Theory
Authors: Fabian Immler and Alexander Maletzky
2016-04-28: No Faster-Than-Light Observers
Authors: Mike Stannett and István Németi
2016-04-27: Algorithms for Reduced Ordered Binary Decision Diagrams
Authors: Julius Michaelis, Maximilian Haslbeck, Peter Lammich and Lars Hupel
2016-04-27: A formalisation of the Cocke-Younger-Kasami algorithm
Author: Maksym Bortin
2016-04-26: Conservation of CSP Noninterference Security under Sequential Composition
Author: Pasquale Noce
2016-04-12: Kleene Algebras with Domain
Authors: Victor B. F. Gomes, Walter Guttmann, Peter Höfner, Georg Struth and Tjark Weber
2016-03-11: Propositional Resolution and Prime Implicates Generation
Author: Nicolas Peltier
2016-03-08: Timed Automata
Author: Simon Wimmer
2016-03-08: The Cartan Fixed Point Theorems
Author: Lawrence C. Paulson
2016-03-01: Linear Temporal Logic
Author: Salomon Sickert
2016-02-17: Analysis of List Update Algorithms
Authors: Maximilian P.L. Haslbeck and Tobias Nipkow
2016-02-05: Verified Construction of Static Single Assignment Form
Authors: Sebastian Ullrich and Denis Lohner
2016-01-29: Polynomial Interpolation
Authors: René Thiemann and Akihisa Yamada
2016-01-29: Polynomial Factorization
Authors: René Thiemann and Akihisa Yamada
2016-01-20: Knot Theory
Author: T.V.H. Prathamesh
2016-01-18: Tensor Product of Matrices
Author: T.V.H. Prathamesh
2016-01-14: Cardinality of Number Partitions
Author: Lukas Bulwahn

 

2015
2015-12-28: Basic Geometric Properties of Triangles
Author: Manuel Eberl
2015-12-28: The Divergence of the Prime Harmonic Series
Author: Manuel Eberl
2015-12-28: Liouville numbers
Author: Manuel Eberl
2015-12-28: Descartes' Rule of Signs
Author: Manuel Eberl
2015-12-22: The Stern-Brocot Tree
Authors: Peter Gammie and Andreas Lochbihler
2015-12-22: Applicative Lifting
Authors: Andreas Lochbihler and Joshua Schneider
2015-12-22: Algebraic Numbers in Isabelle/HOL
Authors: René Thiemann, Akihisa Yamada and Sebastiaan Joosten
2015-12-12: Cardinality of Set Partitions
Author: Lukas Bulwahn
2015-12-02: Latin Square
Author: Alexander Bentkamp
2015-12-01: Ergodic Theory
Author: Sebastien Gouezel
2015-11-19: Euler's Partition Theorem
Author: Lukas Bulwahn
2015-11-18: The Tortoise and Hare Algorithm
Author: Peter Gammie
2015-11-11: Planarity Certificates
Author: Lars Noschinski
2015-11-02: Positional Determinacy of Parity Games
Author: Christoph Dittmann
2015-09-16: A Meta-Model for the Isabelle API
Authors: Frédéric Tuong and Burkhart Wolff
2015-09-04: Converting Linear Temporal Logic to Deterministic (Generalized) Rabin Automata
Author: Salomon Sickert
2015-08-21: Matrices, Jordan Normal Forms, and Spectral Radius Theory
Authors: René Thiemann and Akihisa Yamada
2015-08-20: Decreasing Diagrams II
Author: Bertram Felgenhauer
2015-08-18: The Inductive Unwinding Theorem for CSP Noninterference Security
Author: Pasquale Noce
2015-08-12: Representations of Finite Groups
Author: Jeremy Sylvestre
2015-08-10: Analysing and Comparing Encodability Criteria for Process Calculi
Authors: Kirstin Peters and Rob van Glabbeek
2015-07-21: Generating Cases from Labeled Subgoals
Author: Lars Noschinski
2015-07-14: Landau Symbols
Author: Manuel Eberl
2015-07-14: The Akra-Bazzi theorem and the Master theorem
Author: Manuel Eberl
2015-07-07: Hermite Normal Form
Authors: Jose Divasón and Jesús Aransay
2015-06-27: Derangements Formula
Author: Lukas Bulwahn
2015-06-11: The Ipurge Unwinding Theorem for CSP Noninterference Security
Author: Pasquale Noce
2015-06-11: The Generic Unwinding Theorem for CSP Noninterference Security
Author: Pasquale Noce
2015-06-11: Binary Multirelations
Authors: Hitoshi Furusawa and Georg Struth
2015-06-11: Reasoning about Lists via List Interleaving
Author: Pasquale Noce
2015-06-07: Parameterized Dynamic Tables
Author: Tobias Nipkow
2015-05-28: Derivatives of Logical Formulas
Author: Dmitriy Traytel
2015-05-27: A Zoo of Probabilistic Systems
Authors: Johannes Hölzl, Andreas Lochbihler and Dmitriy Traytel
2015-04-30: VCG - Combinatorial Vickrey-Clarke-Groves Auctions
Authors: Marco B. Caminati, Manfred Kerber, Christoph Lange and Colin Rowat
2015-04-15: Residuated Lattices
Authors: Victor B. F. Gomes and Georg Struth
2015-04-13: Concurrent IMP
Author: Peter Gammie
2015-04-13: Relaxing Safely: Verified On-the-Fly Garbage Collection for x86-TSO
Authors: Peter Gammie, Tony Hosking and Kai Engelhardt
2015-03-30: Trie
Authors: Andreas Lochbihler and Tobias Nipkow
2015-03-18: Consensus Refined
Authors: Ognjen Maric and Christoph Sprenger
2015-03-11: Deriving class instances for datatypes
Authors: Christian Sternagel and René Thiemann
2015-02-20: The Safety of Call Arity
Author: Joachim Breitner
2015-02-12: QR Decomposition
Authors: Jose Divasón and Jesús Aransay
2015-02-12: Echelon Form
Authors: Jose Divasón and Jesús Aransay
2015-02-05: Finite Automata in Hereditarily Finite Set Theory
Author: Lawrence C. Paulson
2015-01-28: Verification of the UpDown Scheme
Author: Johannes Hölzl

 

2014
2014-11-28: The Unified Policy Framework (UPF)
Authors: Achim D. Brucker, Lukas Brügger and Burkhart Wolff
2014-10-23: Loop freedom of the (untimed) AODV routing protocol
Authors: Timothy Bourke and Peter Höfner
2014-10-13: Lifting Definition Option
Author: René Thiemann
2014-10-10: Stream Fusion in HOL with Code Generation
Authors: Andreas Lochbihler and Alexandra Maximova
2014-10-09: A Verified Compiler for Probability Density Functions
Authors: Manuel Eberl, Johannes Hölzl and Tobias Nipkow
2014-10-08: Formalization of Refinement Calculus for Reactive Systems
Author: Viorel Preoteasa
2014-10-03: XML
Authors: Christian Sternagel and René Thiemann
2014-10-03: Certification Monads
Authors: Christian Sternagel and René Thiemann
2014-09-25: Imperative Insertion Sort
Author: Christian Sternagel
2014-09-19: The Sturm-Tarski Theorem
Author: Wenda Li
2014-09-15: The Cayley-Hamilton Theorem
Authors: Stephan Adelsberger, Stefan Hetzl and Florian Pollak
2014-09-09: The Jordan-Hölder Theorem
Author: Jakob von Raumer
2014-09-04: Priority Queues Based on Braun Trees
Author: Tobias Nipkow
2014-09-03: Gauss-Jordan Algorithm and Its Applications
Authors: Jose Divasón and Jesús Aransay
2014-08-29: Vector Spaces
Author: Holden Lee
2014-08-29: Real-Valued Special Functions: Upper and Lower Bounds
Author: Lawrence C. Paulson
2014-08-13: Skew Heap
Author: Tobias Nipkow
2014-08-12: Splay Tree
Author: Tobias Nipkow
2014-07-29: Haskell's Show Class in Isabelle/HOL
Authors: Christian Sternagel and René Thiemann
2014-07-18: Formal Specification of a Generic Separation Kernel
Authors: Freek Verbeek, Sergey Tverdyshev, Oto Havle, Holger Blasum, Bruno Langenstein, Werner Stephan, Yakoub Nemouchi, Abderrahmane Feliachi, Burkhart Wolff and Julien Schmaltz
2014-07-13: pGCL for Isabelle
Author: David Cock
2014-07-07: Amortized Complexity Verified
Author: Tobias Nipkow
2014-07-04: Network Security Policy Verification
Author: Cornelius Diekmann
2014-07-03: Pop-Refinement
Author: Alessandro Coglio
2014-06-12: Decision Procedures for MSO on Words Based on Derivatives of Regular Expressions
Authors: Dmitriy Traytel and Tobias Nipkow
2014-06-08: Boolean Expression Checkers
Author: Tobias Nipkow
2014-05-28: Promela Formalization
Author: René Neumann
2014-05-28: Converting Linear-Time Temporal Logic to Generalized Büchi Automata
Authors: Alexander Schimpf and Peter Lammich
2014-05-28: Verified Efficient Implementation of Gabow's Strongly Connected Components Algorithm
Author: Peter Lammich
2014-05-28: A Fully Verified Executable LTL Model Checker
Authors: Javier Esparza, Peter Lammich, René Neumann, Tobias Nipkow, Alexander Schimpf and Jan-Georg Smaus
2014-05-28: The CAVA Automata Library
Author: Peter Lammich
2014-05-23: Transitive closure according to Roy-Floyd-Warshall
Author: Makarius Wenzel
2014-05-23: Noninterference Security in Communicating Sequential Processes
Author: Pasquale Noce
2014-05-21: Regular Algebras
Authors: Simon Foster and Georg Struth
2014-04-28: Formalisation and Analysis of Component Dependencies
Author: Maria Spichkova
2014-04-23: A Formalization of Declassification with WHAT-and-WHERE-Security
Authors: Sylvia Grewe, Alexander Lux, Heiko Mantel and Jens Sauer
2014-04-23: A Formalization of Strong Security
Authors: Sylvia Grewe, Alexander Lux, Heiko Mantel and Jens Sauer
2014-04-23: A Formalization of Assumptions and Guarantees for Compositional Noninterference
Authors: Sylvia Grewe, Heiko Mantel and Daniel Schoepe
2014-04-22: Bounded-Deducibility Security
Authors: Andrei Popescu, Peter Lammich and Thomas Bauereiss
2014-04-16: A shallow embedding of HyperCTL*
Authors: Markus N. Rabe, Peter Lammich and Andrei Popescu
2014-04-16: Abstract Completeness
Authors: Jasmin Christian Blanchette, Andrei Popescu and Dmitriy Traytel
2014-04-13: Discrete Summation
Author: Florian Haftmann
2014-04-03: Syntax and semantics of a GPU kernel programming language
Author: John Wickerson
2014-03-11: Probabilistic Noninterference
Authors: Andrei Popescu and Johannes Hölzl
2014-03-08: Mechanization of the Algebra for Wireless Networks (AWN)
Author: Timothy Bourke
2014-02-18: Mutually Recursive Partial Functions
Author: René Thiemann
2014-02-13: Properties of Random Graphs -- Subgraph Containment
Author: Lars Hupel
2014-02-11: Verification of Selection and Heap Sort Using Locales
Author: Danijela Petrovic
2014-02-07: Affine Arithmetic
Author: Fabian Immler
2014-02-06: Implementing field extensions of the form Q[sqrt(b)]
Author: René Thiemann
2014-01-30: Unified Decision Procedures for Regular Expression Equivalence
Authors: Tobias Nipkow and Dmitriy Traytel
2014-01-28: Secondary Sylow Theorems
Author: Jakob von Raumer
2014-01-25: Relation Algebra
Authors: Alasdair Armstrong, Simon Foster, Georg Struth and Tjark Weber
2014-01-23: Kleene Algebra with Tests and Demonic Refinement Algebras
Authors: Alasdair Armstrong, Victor B. F. Gomes and Georg Struth
2014-01-16: Featherweight OCL: A Proposal for a Machine-Checked Formal Semantics for OCL 2.5
Authors: Achim D. Brucker, Frédéric Tuong and Burkhart Wolff
2014-01-11: Sturm's Theorem
Author: Manuel Eberl
2014-01-11: Compositional Properties of Crypto-Based Components
Author: Maria Spichkova

 

2013
2013-12-01: A General Method for the Proof of Theorems on Tail-recursive Functions
Author: Pasquale Noce
2013-11-17: Gödel's Incompleteness Theorems
Author: Lawrence C. Paulson
2013-11-17: The Hereditarily Finite Sets
Author: Lawrence C. Paulson
2013-11-15: A Codatatype of Formal Languages
Author: Dmitriy Traytel
2013-11-14: Stream Processing Components: Isabelle/HOL Formalisation and Case Studies
Author: Maria Spichkova
2013-11-12: Gödel's God in Isabelle/HOL
Authors: Christoph Benzmüller and Bruno Woltzenlogel Paleo
2013-11-01: Decreasing Diagrams
Author: Harald Zankl
2013-10-02: Automatic Data Refinement
Author: Peter Lammich
2013-09-17: Native Word
Author: Andreas Lochbihler
2013-07-27: A Formal Model of IEEE Floating Point Arithmetic
Author: Lei Yu
2013-07-22: Pratt's Primality Certificates
Authors: Simon Wimmer and Lars Noschinski
2013-07-22: Lehmer's Theorem
Authors: Simon Wimmer and Lars Noschinski
2013-07-19: The Königsberg Bridge Problem and the Friendship Theorem
Author: Wenda Li
2013-06-27: Sound and Complete Sort Encodings for First-Order Logic
Authors: Jasmin Christian Blanchette and Andrei Popescu
2013-05-22: An Axiomatic Characterization of the Single-Source Shortest Path Problem
Author: Christine Rizkallah
2013-04-28: Graph Theory
Author: Lars Noschinski
2013-04-15: Light-weight Containers
Author: Andreas Lochbihler
2013-02-21: Nominal 2
Authors: Christian Urban, Stefan Berghofer and Cezary Kaliszyk
2013-01-31: The Correctness of Launchbury's Natural Semantics for Lazy Evaluation
Author: Joachim Breitner
2013-01-19: Ribbon Proofs
Author: John Wickerson
2013-01-16: Rank-Nullity Theorem in Linear Algebra
Authors: Jose Divasón and Jesús Aransay
2013-01-15: Kleene Algebra
Authors: Alasdair Armstrong, Georg Struth and Tjark Weber
2013-01-03: Computing N-th Roots using the Babylonian Method
Author: René Thiemann

 

2012
2012-11-14: A Separation Logic Framework for Imperative HOL
Authors: Peter Lammich and Rene Meis
2012-11-02: Open Induction
Authors: Mizuhito Ogawa and Christian Sternagel
2012-10-30: The independence of Tarski's Euclidean axiom
Author: T. J. M. Makarios
2012-10-27: Bondy's Theorem
Authors: Jeremy Avigad and Stefan Hetzl
2012-09-10: Possibilistic Noninterference
Authors: Andrei Popescu and Johannes Hölzl
2012-08-07: Generating linear orders for datatypes
Author: René Thiemann
2012-08-05: Proving the Impossibility of Trisecting an Angle and Doubling the Cube
Authors: Ralph Romanos and Lawrence C. Paulson
2012-07-27: Verifying Fault-Tolerant Distributed Algorithms in the Heard-Of Model
Authors: Henri Debrat and Stephan Merz
2012-07-01: Logical Relations for PCF
Author: Peter Gammie
2012-06-26: Type Constructor Classes and Monad Transformers
Author: Brian Huffman
2012-05-29: Psi-calculi in Isabelle
Author: Jesper Bengtson
2012-05-29: The pi-calculus in nominal logic
Author: Jesper Bengtson
2012-05-29: CCS in nominal logic
Author: Jesper Bengtson
2012-05-27: Isabelle/Circus
Authors: Abderrahmane Feliachi, Burkhart Wolff and Marie-Claude Gaudel
2012-05-11: Separation Algebra
Authors: Gerwin Klein, Rafal Kolanski and Andrew Boyton
2012-05-07: Stuttering Equivalence
Author: Stephan Merz
2012-05-02: Inductive Study of Confidentiality
Author: Giampaolo Bella
2012-04-26: Ordinary Differential Equations
Authors: Fabian Immler and Johannes Hölzl
2012-04-13: Well-Quasi-Orders
Author: Christian Sternagel
2012-03-01: Abortable Linearizable Modules
Authors: Rachid Guerraoui, Viktor Kuncak and Giuliano Losa
2012-02-29: Executable Transitive Closures
Author: René Thiemann
2012-02-06: A Probabilistic Proof of the Girth-Chromatic Number Theorem
Author: Lars Noschinski
2012-01-30: Refinement for Monadic Programs
Author: Peter Lammich
2012-01-30: Dijkstra's Shortest Path Algorithm
Authors: Benedikt Nordhoff and Peter Lammich
2012-01-03: Markov Models
Authors: Johannes Hölzl and Tobias Nipkow

 

2011
2011-11-19: A Definitional Encoding of TLA* in Isabelle/HOL
Authors: Gudmund Grov and Stephan Merz
2011-11-09: Efficient Mergesort
Author: Christian Sternagel
2011-09-22: Pseudo Hoops
Authors: George Georgescu, Laurentiu Leustean and Viorel Preoteasa
2011-09-22: Algebra of Monotonic Boolean Transformers
Author: Viorel Preoteasa
2011-09-22: Lattice Properties
Author: Viorel Preoteasa
2011-08-26: The Myhill-Nerode Theorem Based on Regular Expressions
Authors: Chunhan Wu, Xingyuan Zhang and Christian Urban
2011-08-19: Gauss-Jordan Elimination for Matrices Represented as Functions
Author: Tobias Nipkow
2011-07-21: Maximum Cardinality Matching
Author: Christine Rizkallah
2011-05-17: Knowledge-based programs
Author: Peter Gammie
2011-04-01: The General Triangle Is Unique
Author: Joachim Breitner
2011-03-14: Executable Transitive Closures of Finite Relations
Authors: Christian Sternagel and René Thiemann
2011-02-23: Interval Temporal Logic on Natural Numbers
Author: David Trachtenherz
2011-02-23: Infinite Lists
Author: David Trachtenherz
2011-02-23: AutoFocus Stream Processing for Single-Clocking and Multi-Clocking Semantics
Author: David Trachtenherz
2011-02-07: Lightweight Java
Authors: Rok Strniša and Matthew Parkinson
2011-01-10: RIPEMD-160
Author: Fabian Immler
2011-01-08: Lower Semicontinuous Functions
Author: Bogdan Grechuk

 

2010
2010-12-17: Hall's Marriage Theorem
Authors: Dongchen Jiang and Tobias Nipkow
2010-11-16: Shivers' Control Flow Analysis
Author: Joachim Breitner
2010-10-28: Finger Trees
Authors: Benedikt Nordhoff, Stefan Körner and Peter Lammich
2010-10-28: Functional Binomial Queues
Author: René Neumann
2010-10-28: Binomial Heaps and Skew Binomial Heaps
Authors: Rene Meis, Finn Nielsen and Peter Lammich
2010-08-29: Strong Normalization of Moggis's Computational Metalanguage
Author: Christian Doczkal
2010-08-10: Executable Multivariate Polynomials
Authors: Christian Sternagel, René Thiemann, Alexander Maletzky, Fabian Immler, Florian Haftmann, Andreas Lochbihler and Alexander Bentkamp
2010-08-08: Formalizing Statecharts using Hierarchical Automata
Authors: Steffen Helke and Florian Kammüller
2010-06-24: Free Groups
Author: Joachim Breitner
2010-06-20: Category Theory
Author: Alexander Katovsky
2010-06-17: Executable Matrix Operations on Matrices of Arbitrary Dimensions
Authors: Christian Sternagel and René Thiemann
2010-06-14: Abstract Rewriting
Authors: Christian Sternagel and René Thiemann
2010-05-28: Verification of the Deutsch-Schorr-Waite Graph Marking Algorithm using Data Refinement
Authors: Viorel Preoteasa and Ralph-Johan Back
2010-05-28: Semantics and Data Refinement of Invariant Based Programs
Authors: Viorel Preoteasa and Ralph-Johan Back
2010-05-22: A Complete Proof of the Robbins Conjecture
Author: Matthew Wampler-Doty
2010-05-12: Regular Sets and Expressions
Authors: Alexander Krauss and Tobias Nipkow
2010-04-30: Locally Nameless Sigma Calculus
Authors: Ludovic Henrio, Florian Kammüller, Bianca Lutz and Henry Sudhof
2010-03-29: Free Boolean Algebra
Author: Brian Huffman
2010-03-23: Inter-Procedural Information Flow Noninterference via Slicing
Author: Daniel Wasserrab
2010-03-23: Information Flow Noninterference via Slicing
Author: Daniel Wasserrab
2010-02-20: List Index
Author: Tobias Nipkow
2010-02-12: Coinductive
Author: Andreas Lochbihler

 

2009
2009-12-09: A Fast SAT Solver for Isabelle in Standard ML
Author: Armin Heller
2009-12-03: Formalizing the Logic-Automaton Connection
Authors: Stefan Berghofer and Markus Reiter
2009-11-25: Tree Automata
Author: Peter Lammich
2009-11-25: Collections Framework
Author: Peter Lammich
2009-11-22: Perfect Number Theorem
Author: Mark Ijbema
2009-11-13: Backing up Slicing: Verifying the Interprocedural Two-Phase Horwitz-Reps-Binkley Slicer
Author: Daniel Wasserrab
2009-10-30: The Worker/Wrapper Transformation
Author: Peter Gammie
2009-09-01: Ordinals and Cardinals
Author: Andrei Popescu
2009-08-28: Invertibility in Sequent Calculi
Author: Peter Chapman
2009-08-04: An Example of a Cofinitary Group in Isabelle/HOL
Author: Bart Kastermans
2009-05-06: Code Generation for Functions as Data
Author: Andreas Lochbihler
2009-04-29: Stream Fusion
Author: Brian Huffman

 

2008
2008-12-12: A Bytecode Logic for JML and Types
Authors: Lennart Beringer and Martin Hofmann
2008-11-10: Secure information flow and program logics
Authors: Lennart Beringer and Martin Hofmann
2008-11-09: Some classical results in Social Choice Theory
Author: Peter Gammie
2008-11-07: Fun With Tilings
Authors: Tobias Nipkow and Lawrence C. Paulson
2008-10-15: The Textbook Proof of Huffman's Algorithm
Author: Jasmin Christian Blanchette
2008-09-16: Towards Certified Slicing
Author: Daniel Wasserrab
2008-09-02: A Correctness Proof for the Volpano/Smith Security Typing System
Authors: Gregor Snelting and Daniel Wasserrab
2008-09-01: Arrow and Gibbard-Satterthwaite
Author: Tobias Nipkow
2008-08-26: Fun With Functions
Author: Tobias Nipkow
2008-07-23: Formal Verification of Modern SAT Solvers
Author: Filip Marić
2008-04-05: Recursion Theory I
Author: Michael Nedzelsky
2008-02-29: A Sequential Imperative Programming Language Syntax, Semantics, Hoare Logics and Verification Environment
Author: Norbert Schirmer
2008-02-29: BDD Normalisation
Authors: Veronika Ortner and Norbert Schirmer
2008-02-18: Normalization by Evaluation
Authors: Klaus Aehlig and Tobias Nipkow
2008-01-11: Quantifier Elimination for Linear Arithmetic
Author: Tobias Nipkow

 

2007
2007-12-14: Formalization of Conflict Analysis of Programs with Procedures, Thread Creation, and Monitors
Authors: Peter Lammich and Markus Müller-Olm
2007-12-03: Jinja with Threads
Author: Andreas Lochbihler
2007-11-06: Much Ado About Two
Author: Sascha Böhme
2007-08-12: Sums of Two and Four Squares
Author: Roelof Oosterhuis
2007-08-12: Fermat's Last Theorem for Exponents 3 and 4 and the Parametrisation of Pythagorean Triples
Author: Roelof Oosterhuis
2007-08-08: Fundamental Properties of Valuation Theory and Hensel's Lemma
Author: Hidetsune Kobayashi
2007-08-02: POPLmark Challenge Via de Bruijn Indices
Author: Stefan Berghofer
2007-08-02: First-Order Logic According to Fitting
Author: Stefan Berghofer

 

2006
2006-09-09: Hotel Key Card System
Author: Tobias Nipkow
2006-08-08: Abstract Hoare Logics
Author: Tobias Nipkow
2006-05-22: Flyspeck I: Tame Graphs
Authors: Gertrud Bauer and Tobias Nipkow
2006-05-15: CoreC++
Author: Daniel Wasserrab
2006-03-31: A Theory of Featherweight Java in Isabelle/HOL
Authors: J. Nathan Foster and Dimitrios Vytiniotis
2006-03-15: Instances of Schneider's generalized protocol of clock synchronization
Author: Damián Barsotti
2006-03-14: Cauchy's Mean Theorem and the Cauchy-Schwarz Inequality
Author: Benjamin Porter

 

2005
2005-11-11: Countable Ordinals
Author: Brian Huffman
2005-10-12: Fast Fourier Transform
Author: Clemens Ballarin
2005-06-24: Formalization of a Generalized Protocol for Clock Synchronization
Author: Alwen Tiu
2005-06-22: Proving the Correctness of Disk Paxos
Authors: Mauro Jaskelioff and Stephan Merz
2005-06-20: Jive Data and Store Model
Authors: Nicole Rauch and Norbert Schirmer
2005-06-01: Jinja is not Java
Authors: Gerwin Klein and Tobias Nipkow
2005-05-02: SHA1, RSA, PSS and more
Authors: Christina Lindenberg and Kai Wirt
2005-04-21: Category Theory to Yoneda's Lemma
Author: Greg O'Keefe

 

2004
2004-12-09: File Refinement
Authors: Karen Zee and Viktor Kuncak
2004-11-19: Integration theory and random variables
Author: Stefan Richter
2004-09-28: A Mechanically Verified, Efficient, Sound and Complete Theorem Prover For First Order Logic
Author: Tom Ridge
2004-09-20: Ramsey's theorem, infinitary version
Author: Tom Ridge
2004-09-20: Completeness theorem
Authors: James Margetson and Tom Ridge
2004-07-09: Compiling Exceptions Correctly
Author: Tobias Nipkow
2004-06-24: Depth First Search
Authors: Toshiaki Nishihara and Yasuhiko Minamide
2004-05-18: Groups, Rings and Modules
Authors: Hidetsune Kobayashi, L. Chen and H. Murao
2004-04-26: Topology
Author: Stefan Friedrich
2004-04-26: Lazy Lists II
Author: Stefan Friedrich
2004-04-05: Binary Search Trees
Author: Viktor Kuncak
2004-03-30: Functional Automata
Author: Tobias Nipkow
2004-03-19: Mini ML
Authors: Wolfgang Naraschewski and Tobias Nipkow
2004-03-19: AVL Trees
Authors: Tobias Nipkow and Cornelia Pusch
\ No newline at end of file diff --git a/web/rss.xml b/web/rss.xml --- a/web/rss.xml +++ b/web/rss.xml @@ -1,564 +1,595 @@ Archive of Formal Proofs https://www.isa-afp.org The Archive of Formal Proofs is a collection of proof libraries, examples, and larger scientific developments, mechanically checked in the theorem prover Isabelle. - 28 Oct 2021 00:00:00 +0000 + 08 Nov 2021 00:00:00 +0000 + + Real Exponents as the Limits of Sequences of Rational Exponents + https://www.isa-afp.org/entries/Real_Power.html + https://www.isa-afp.org/entries/Real_Power.html + Jacques D. Fleuriot + 08 Nov 2021 00:00:00 +0000 + +In this formalisation, we construct real exponents as the limits of +sequences of rational exponents. In particular, if $a \ge 1$ and $x +\in \mathbb{R}$, we choose an increasing rational sequence $r_n$ such +that $\lim_{n\to\infty} {r_n} = x$. Then the sequence $a^{r_n}$ is +increasing and if $r$ is any rational number such that $r > x$, +$a^{r_n}$ is bounded above by $a^r$. By the convergence criterion for +monotone sequences, $a^{r_n}$ converges. We define $a^ x = +\lim_{n\to\infty} a^{r_n}$ and show that it has the expected +properties (for $a \ge 0$). This particular construction of real +exponents is needed instead of the usual one using the natural +logarithm and exponential functions (which already exists in Isabelle) +to support our mechanical derivation of Euler's exponential +series as an ``infinite polynomial". Aside from helping us avoid +circular reasoning, this is, as far as we are aware, the first time +real exponents are mechanised in this way within a proof assistant. + + + Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL + https://www.isa-afp.org/entries/PAL.html + https://www.isa-afp.org/entries/PAL.html + Christoph Benzmüller, Sebastian Reiche + 08 Nov 2021 00:00:00 +0000 + +We present a shallow embedding of public announcement logic (PAL) with +relativized general knowledge in HOL. We then use PAL to obtain an +elegant encoding of the wise men puzzle, which we solve automatically +using sledgehammer. + + + Factorization of Polynomials with Algebraic Coefficients + https://www.isa-afp.org/entries/Factor_Algebraic_Polynomial.html + https://www.isa-afp.org/entries/Factor_Algebraic_Polynomial.html + Manuel Eberl, René Thiemann + 08 Nov 2021 00:00:00 +0000 + +The AFP already contains a verified implementation of algebraic +numbers. However, it is has a severe limitation in its factorization +algorithm of real and complex polynomials: the factorization is only +guaranteed to succeed if the coefficients of the polynomial are +rational numbers. In this work, we verify an algorithm to factor all +real and complex polynomials whose coefficients are algebraic. The +existence of such an algorithm proves in a constructive way that the +set of complex algebraic numbers is algebraically closed. Internally, +the algorithm is based on resultants of multivariate polynomials and +an approximation algorithm using interval arithmetic. + + + Szemerédi's Regularity Lemma + https://www.isa-afp.org/entries/Szemeredi_Regularity.html + https://www.isa-afp.org/entries/Szemeredi_Regularity.html + Chelsea Edmonds, Angeliki Koutsoukou-Argyraki, Lawrence C. Paulson + 05 Nov 2021 00:00:00 +0000 + +<a +href="https://en.wikipedia.org/wiki/Szemerédi_regularity_lemma">Szemerédi's +regularity lemma</a> is a key result in the study of large +graphs. It asserts the existence an upper bound on the number of parts +the vertices of a graph need to be partitioned into such that the +edges between the parts are random in a certain sense. This bound +depends only on the desired precision and not on the graph itself, in +the spirit of Ramsey's theorem. The formalisation follows online +course notes by <a +href="https://www.dpmms.cam.ac.uk/~par31/notes/tic.pdf">Tim +Gowers</a> and <a +href="https://yufeizhao.com/gtac/gtac.pdf">Yufei +Zhao</a>. + Quantum and Classical Registers https://www.isa-afp.org/entries/Registers.html https://www.isa-afp.org/entries/Registers.html Dominique Unruh 28 Oct 2021 00:00:00 +0000 A formalization of the theory of quantum and classical registers as developed by (Unruh, Quantum and Classical Registers). In a nutshell, a register refers to a part of a larger memory or system that can be accessed independently. Registers can be constructed from other registers and several (compatible) registers can be composed. This formalization develops both the generic theory of registers as well as specific instantiations for classical and quantum registers. Belief Revision Theory https://www.isa-afp.org/entries/Belief_Revision.html https://www.isa-afp.org/entries/Belief_Revision.html Valentin Fouillard, Safouan Taha, Frédéric Boulanger, Nicolas Sabouret 19 Oct 2021 00:00:00 +0000 The 1985 paper by Carlos Alchourrón, Peter Gärdenfors, and David Makinson (AGM), “On the Logic of Theory Change: Partial Meet Contraction and Revision Functions” launches a large and rapidly growing literature that employs formal models and logics to handle changing beliefs of a rational agent and to take into account new piece of information observed by this agent. In 2011, a review book titled "AGM 25 Years: Twenty-Five Years of Research in Belief Change" was edited to summarize the first twenty five years of works based on AGM. This HOL-based AFP entry is a faithful formalization of the AGM operators (e.g. contraction, revision, remainder ...) axiomatized in the original paper. It also contains the proofs of all the theorems stated in the paper that show how these operators combine. Both proofs of Harper and Levi identities are established. X86 instruction semantics and basic block symbolic execution https://www.isa-afp.org/entries/X86_Semantics.html https://www.isa-afp.org/entries/X86_Semantics.html Freek Verbeek, Abhijith Bharadwaj, Joshua Bockenek, Ian Roessle, Timmy Weerwag, Binoy Ravindran 13 Oct 2021 00:00:00 +0000 This AFP entry provides semantics for roughly 120 different X86-64 assembly instructions. These instructions include various moves, arithmetic/logical operations, jumps, call/return, SIMD extensions and others. External functions are supported by allowing a user to provide custom semantics for these calls. Floating-point operations are mapped to uninterpreted functions. The model provides semantics for register aliasing and a byte-level little-endian memory model. The semantics are purposefully incomplete, but overapproximative. For example, the precise effect of flags may be undefined for certain instructions, or instructions may simply have no semantics at all. In those cases, the semantics are mapped to universally quantified uninterpreted terms from a locale. Second, this entry provides a method to symbolic execution of basic blocks. The method, called ''se_step'' (for: symbolic execution step) fetches an instruction and updates the current symbolic state while keeping track of assumptions made over the memory model. A key component is a set of theorems that prove how reads from memory resolve after writes have occurred. Thirdly, this entry provides a parser that allows the user to copy-paste the output of the standard disassembly tool objdump into Isabelle/HOL. A couple small and explanatory examples are included, including functions from the word count program. Several examples can be supplied upon request (they are not included due to the running time of verification): functions from the floating-point modulo function from FDLIBM, the GLIBC strlen function and the CoreUtils SHA256 implementation. Algebras for Iteration, Infinite Executions and Correctness of Sequential Computations https://www.isa-afp.org/entries/Correctness_Algebras.html https://www.isa-afp.org/entries/Correctness_Algebras.html Walter Guttmann 12 Oct 2021 00:00:00 +0000 We study models of state-based non-deterministic sequential computations and describe them using algebras. We propose algebras that describe iteration for strict and non-strict computations. They unify computation models which differ in the fixpoints used to represent iteration. We propose algebras that describe the infinite executions of a computation. They lead to a unified approximation order and results that connect fixpoints in the approximation and refinement orders. This unifies the semantics of recursion for a range of computation models. We propose algebras that describe preconditions and the effect of while-programs under postconditions. They unify correctness statements in two dimensions: one statement applies in various computation models to various correctness claims. Verified Quadratic Virtual Substitution for Real Arithmetic https://www.isa-afp.org/entries/Virtual_Substitution.html https://www.isa-afp.org/entries/Virtual_Substitution.html Matias Scharager, Katherine Cordwell, Stefan Mitsch, André Platzer 02 Oct 2021 00:00:00 +0000 This paper presents a formally verified quantifier elimination (QE) algorithm for first-order real arithmetic by linear and quadratic virtual substitution (VS) in Isabelle/HOL. The Tarski-Seidenberg theorem established that the first-order logic of real arithmetic is decidable by QE. However, in practice, QE algorithms are highly complicated and often combine multiple methods for performance. VS is a practically successful method for QE that targets formulas with low-degree polynomials. To our knowledge, this is the first work to formalize VS for quadratic real arithmetic including inequalities. The proofs necessitate various contributions to the existing multivariate polynomial libraries in Isabelle/HOL. Our framework is modularized and easily expandable (to facilitate integrating future optimizations), and could serve as a basis for developing practical general-purpose QE algorithms. Further, as our formalization is designed with practicality in mind, we export our development to SML and test the resulting code on 378 benchmarks from the literature, comparing to Redlog, Z3, Wolfram Engine, and SMT-RAT. This identified inconsistencies in some tools, underscoring the significance of a verified approach for the intricacies of real arithmetic. Soundness and Completeness of an Axiomatic System for First-Order Logic https://www.isa-afp.org/entries/FOL_Axiomatic.html https://www.isa-afp.org/entries/FOL_Axiomatic.html Asta Halkjær From 24 Sep 2021 00:00:00 +0000 This work is a formalization of the soundness and completeness of an axiomatic system for first-order logic. The proof system is based on System Q1 by Smullyan and the completeness proof follows his textbook "First-Order Logic" (Springer-Verlag 1968). The completeness proof is in the Henkin style where a consistent set is extended to a maximal consistent set using Lindenbaum's construction and Henkin witnesses are added during the construction to ensure saturation as well. The resulting set is a Hintikka set which, by the model existence theorem, is satisfiable in the Herbrand universe. Complex Bounded Operators https://www.isa-afp.org/entries/Complex_Bounded_Operators.html https://www.isa-afp.org/entries/Complex_Bounded_Operators.html Jose Manuel Rodriguez Caballero, Dominique Unruh 18 Sep 2021 00:00:00 +0000 We present a formalization of bounded operators on complex vector spaces. Our formalization contains material on complex vector spaces (normed spaces, Banach spaces, Hilbert spaces) that complements and goes beyond the developments of real vectors spaces in the Isabelle/HOL standard library. We define the type of bounded operators between complex vector spaces (<em>cblinfun</em>) and develop the theory of unitaries, projectors, extension of bounded linear functions (BLT theorem), adjoints, Loewner order, closed subspaces and more. For the finite-dimensional case, we provide code generation support by identifying finite-dimensional operators with matrices as formalized in the <a href="Jordan_Normal_Form.html">Jordan_Normal_Form</a> AFP entry. A Formalization of Weighted Path Orders and Recursive Path Orders https://www.isa-afp.org/entries/Weighted_Path_Order.html https://www.isa-afp.org/entries/Weighted_Path_Order.html Christian Sternagel, René Thiemann, Akihisa Yamada 16 Sep 2021 00:00:00 +0000 We define the weighted path order (WPO) and formalize several properties such as strong normalization, the subterm property, and closure properties under substitutions and contexts. Our definition of WPO extends the original definition by also permitting multiset comparisons of arguments instead of just lexicographic extensions. Therefore, our WPO not only subsumes lexicographic path orders (LPO), but also recursive path orders (RPO). We formally prove these subsumptions and therefore all of the mentioned properties of WPO are automatically transferable to LPO and RPO as well. Such a transformation is not required for Knuth&ndash;Bendix orders (KBO), since they have already been formalized. Nevertheless, we still provide a proof that WPO subsumes KBO and thereby underline the generality of WPO. Extension of Types-To-Sets https://www.isa-afp.org/entries/Types_To_Sets_Extension.html https://www.isa-afp.org/entries/Types_To_Sets_Extension.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 In their article titled <i>From Types to Sets by Local Type Definitions in Higher-Order Logic</i> and published in the proceedings of the conference <i>Interactive Theorem Proving</i> in 2016, Ondřej Kunčar and Andrei Popescu propose an extension of the logic Isabelle/HOL and an associated algorithm for the relativization of the <i>type-based theorems</i> to more flexible <i>set-based theorems</i>, collectively referred to as <i>Types-To-Sets</i>. One of the aims of their work was to open an opportunity for the development of a software tool for applied relativization in the implementation of the logic Isabelle/HOL of the proof assistant Isabelle. In this article, we provide a prototype of a software framework for the interactive automated relativization of theorems in Isabelle/HOL, developed as an extension of the proof language Isabelle/Isar. The software framework incorporates the implementation of the proposed extension of the logic, and builds upon some of the ideas for further work expressed in the original article on Types-To-Sets by Ondřej Kunčar and Andrei Popescu and the subsequent article <i>Smooth Manifolds and Types to Sets for Linear Algebra in Isabelle/HOL</i> that was written by Fabian Immler and Bohua Zhan and published in the proceedings of the <i>International Conference on Certified Programs and Proofs</i> in 2019. IDE: Introduction, Destruction, Elimination https://www.isa-afp.org/entries/Intro_Dest_Elim.html https://www.isa-afp.org/entries/Intro_Dest_Elim.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 The article provides the command <b>mk_ide</b> for the object logic Isabelle/HOL of the formal proof assistant Isabelle. The command <b>mk_ide</b> enables the automated synthesis of the introduction, destruction and elimination rules from arbitrary definitions of constant predicates stated in Isabelle/HOL. Conditional Transfer Rule https://www.isa-afp.org/entries/Conditional_Transfer_Rule.html https://www.isa-afp.org/entries/Conditional_Transfer_Rule.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 This article provides a collection of experimental utilities for unoverloading of definitions and synthesis of conditional transfer rules for the object logic Isabelle/HOL of the formal proof assistant Isabelle written in Isabelle/ML. Conditional Simplification https://www.isa-afp.org/entries/Conditional_Simplification.html https://www.isa-afp.org/entries/Conditional_Simplification.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 The article provides a collection of experimental general-purpose proof methods for the object logic Isabelle/HOL of the formal proof assistant Isabelle. The methods in the collection offer functionality that is similar to certain aspects of the functionality provided by the standard proof methods of Isabelle that combine classical reasoning and rewriting, such as the method <i>auto</i>, but use a different approach for rewriting. More specifically, these methods allow for the side conditions of the rewrite rules to be solved via intro-resolution. Category Theory for ZFC in HOL III: Universal Constructions https://www.isa-afp.org/entries/CZH_Universal_Constructions.html https://www.isa-afp.org/entries/CZH_Universal_Constructions.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 The article provides a formalization of elements of the theory of universal constructions for 1-categories (such as limits, adjoints and Kan extensions) in the object logic ZFC in HOL of the formal proof assistant Isabelle. The article builds upon the foundations established in the AFP entry <i>Category Theory for ZFC in HOL II: Elementary Theory of 1-Categories</i>. Category Theory for ZFC in HOL I: Foundations: Design Patterns, Set Theory, Digraphs, Semicategories https://www.isa-afp.org/entries/CZH_Foundations.html https://www.isa-afp.org/entries/CZH_Foundations.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 This article provides a foundational framework for the formalization of category theory in the object logic ZFC in HOL of the formal proof assistant Isabelle. More specifically, this article provides a formalization of canonical set-theoretic constructions internalized in the type <i>V</i> associated with the ZFC in HOL, establishes a design pattern for the formalization of mathematical structures using sequences and locales, and showcases the developed infrastructure by providing formalizations of the elementary theories of digraphs and semicategories. The methodology chosen for the formalization of the theories of digraphs and semicategories (and categories in future articles) rests on the ideas that were originally expressed in the article <i>Set-Theoretical Foundations of Category Theory</i> written by Solomon Feferman and Georg Kreisel. Thus, in the context of this work, each of the aforementioned mathematical structures is represented as a term of the type <i>V</i> embedded into a stage of the von Neumann hierarchy. Category Theory for ZFC in HOL II: Elementary Theory of 1-Categories https://www.isa-afp.org/entries/CZH_Elementary_Categories.html https://www.isa-afp.org/entries/CZH_Elementary_Categories.html Mihails Milehins 06 Sep 2021 00:00:00 +0000 This article provides a formalization of the foundations of the theory of 1-categories in the object logic ZFC in HOL of the formal proof assistant Isabelle. The article builds upon the foundations that were established in the AFP entry <i>Category Theory for ZFC in HOL I: Foundations: Design Patterns, Set Theory, Digraphs, Semicategories</i>. A data flow analysis algorithm for computing dominators https://www.isa-afp.org/entries/Dominance_CHK.html https://www.isa-afp.org/entries/Dominance_CHK.html Nan Jiang 05 Sep 2021 00:00:00 +0000 This entry formalises the fast iterative algorithm for computing dominators due to Cooper, Harvey and Kennedy. It gives a specification of computing dominators on a control flow graph where each node refers to its reverse post order number. A semilattice of reversed-ordered list which represents dominators is built and a Kildall-style algorithm on the semilattice is defined for computing dominators. Finally the soundness and completeness of the algorithm are proved w.r.t. the specification. Solving Cubic and Quartic Equations https://www.isa-afp.org/entries/Cubic_Quartic_Equations.html https://www.isa-afp.org/entries/Cubic_Quartic_Equations.html René Thiemann 03 Sep 2021 00:00:00 +0000 <p>We formalize Cardano's formula to solve a cubic equation $$ax^3 + bx^2 + cx + d = 0,$$ as well as Ferrari's formula to solve a quartic equation. We further turn both formulas into executable algorithms based on the algebraic number implementation in the AFP. To this end we also slightly extended this library, namely by making the minimal polynomial of an algebraic number executable, and by defining and implementing $n$-th roots of complex numbers.</p> Logging-independent Message Anonymity in the Relational Method https://www.isa-afp.org/entries/Logging_Independent_Anonymity.html https://www.isa-afp.org/entries/Logging_Independent_Anonymity.html Pasquale Noce 26 Aug 2021 00:00:00 +0000 In the context of formal cryptographic protocol verification, logging-independent message anonymity is the property for a given message to remain anonymous despite the attacker's capability of mapping messages of that sort to agents based on some intrinsic feature of such messages, rather than by logging the messages exchanged by legitimate agents as with logging-dependent message anonymity. This paper illustrates how logging-independent message anonymity can be formalized according to the relational method for formal protocol verification by considering a real-world protocol, namely the Restricted Identification one by the BSI. This sample model is used to verify that the pseudonymous identifiers output by user identification tokens remain anonymous under the expected conditions. The Theorem of Three Circles https://www.isa-afp.org/entries/Three_Circles.html https://www.isa-afp.org/entries/Three_Circles.html Fox Thomson, Wenda Li 21 Aug 2021 00:00:00 +0000 The Descartes test based on Bernstein coefficients and Descartes’ rule of signs effectively (over-)approximates the number of real roots of a univariate polynomial over an interval. In this entry we formalise the theorem of three circles, which gives sufficient conditions for when the Descartes test returns 0 or 1. This is the first step for efficient root isolation. Fresh identifiers https://www.isa-afp.org/entries/Fresh_Identifiers.html https://www.isa-afp.org/entries/Fresh_Identifiers.html Andrei Popescu, Thomas Bauereiss 16 Aug 2021 00:00:00 +0000 This entry defines a type class with an operator returning a fresh identifier, given a set of already used identifiers and a preferred identifier. The entry provides a default instantiation for any infinite type, as well as executable instantiations for natural numbers and strings. + CoSMed: A confidentiality-verified social media platform + https://www.isa-afp.org/entries/CoSMed.html + https://www.isa-afp.org/entries/CoSMed.html + Thomas Bauereiss, Andrei Popescu + 16 Aug 2021 00:00:00 +0000 + +This entry contains the confidentiality verification of the +(functional kernel of) the CoSMed social media platform. The +confidentiality properties are formalized as instances of BD Security +[<a href="https://doi.org/10.4230/LIPIcs.ITP.2021.3">1</a>, +<a href="https://www.isa-afp.org/entries/Bounded_Deducibility_Security.html">2</a>]. +An innovation in the deployment of BD Security compared to previous +work is the use of dynamic declassification triggers, incorporated as +part of inductive bounds, for providing stronger guarantees that +account for the repeated opening and closing of access windows. To +further strengthen the confidentiality guarantees, we also prove +"traceback" properties about the accessibility decisions +affecting the information managed by the system. + + + CoSMeDis: A confidentiality-verified distributed social media platform + https://www.isa-afp.org/entries/CoSMeDis.html + https://www.isa-afp.org/entries/CoSMeDis.html + Thomas Bauereiss, Andrei Popescu + 16 Aug 2021 00:00:00 +0000 + +This entry contains the confidentiality verification of the +(functional kernel of) the CoSMeDis distributed social media platform +presented in [<a href="https://doi.org/10.1109/SP.2017.24">1</a>]. +CoSMeDis is a multi-node extension the CoSMed prototype social media +platform [<a href="https://doi.org/10.1007/978-3-319-43144-4_6">2</a>, +<a href="https://doi.org/10.1007/s10817-017-9443-3">3</a>, +<a href="https://www.isa-afp.org/entries/CoSMed.html">4</a>]. +The confidentiality properties are formalized as instances of BD +Security [<a href="https://doi.org/10.4230/LIPIcs.ITP.2021.3">5</a>, +<a href="https://www.isa-afp.org/entries/Bounded_Deducibility_Security.html">6</a>]. +The lifting of confidentiality properties from single nodes to the +entire CoSMeDis network is performed using compositionality and +transport theorems for BD Security, which are described in [<a +href="https://doi.org/10.1109/SP.2017.24">1</a>] +and formalized in a separate <a +href="https://www.isa-afp.org/entries/BD_Security_Compositional.html">AFP +entry</a>. + + + CoCon: A Confidentiality-Verified Conference Management System + https://www.isa-afp.org/entries/CoCon.html + https://www.isa-afp.org/entries/CoCon.html + Andrei Popescu, Peter Lammich, Thomas Bauereiss + 16 Aug 2021 00:00:00 +0000 + +This entry contains the confidentiality verification of the +(functional kernel of) the CoCon conference management system [<a +href="https://doi.org/10.1007/978-3-319-08867-9_11">1</a>, +<a href="https://doi.org/10.1007/s10817-020-09566-9">2</a>]. +The confidentiality properties refer to the documents managed by the +system, namely papers, reviews, discussion logs and +acceptance/rejection decisions, and also to the assignment of +reviewers to papers. They have all been formulated as instances of BD +Security [<a href="https://doi.org/10.4230/LIPIcs.ITP.2021.3">3</a>, +<a href="https://www.isa-afp.org/entries/Bounded_Deducibility_Security.html">4</a>] +and verified using the BD Security unwinding technique. + + + Compositional BD Security + https://www.isa-afp.org/entries/BD_Security_Compositional.html + https://www.isa-afp.org/entries/BD_Security_Compositional.html + Thomas Bauereiss, Andrei Popescu + 16 Aug 2021 00:00:00 +0000 + +Building on a previous <a +href="https://www.isa-afp.org/entries/Bounded_Deducibility_Security.html">AFP +entry</a> that formalizes the Bounded-Deducibility Security (BD +Security) framework <a href="https://doi.org/10.4230/LIPIcs.ITP.2021.3">[1]</a>, +we formalize compositionality and transport theorems for information +flow security. These results allow lifting BD Security properties from +individual components specified as transition systems, to a +composition of systems specified as communicating products of +transition systems. The underlying ideas of these results are +presented in the papers <a href="https://doi.org/10.4230/LIPIcs.ITP.2021.3">[1]</a> +and <a href="https://doi.org/10.1109/SP.2017.24">[2]</a>. +The latter paper also describes a major case study where these results +have been used: on verifying the CoSMeDis distributed social media +platform (itself formalized as an <a +href="https://www.isa-afp.org/entries/CoSMeDis.html">AFP +entry</a> that builds on this entry). + + Combinatorial Design Theory https://www.isa-afp.org/entries/Design_Theory.html https://www.isa-afp.org/entries/Design_Theory.html Chelsea Edmonds, Lawrence Paulson 13 Aug 2021 00:00:00 +0000 Combinatorial design theory studies incidence set systems with certain balance and symmetry properties. It is closely related to hypergraph theory. This formalisation presents a general library for formal reasoning on incidence set systems, designs and their applications, including formal definitions and proofs for many key properties, operations, and theorems on the construction and existence of designs. Notably, this includes formalising t-designs, balanced incomplete block designs (BIBD), group divisible designs (GDD), pairwise balanced designs (PBD), design isomorphisms, and the relationship between graphs and designs. A locale-centric approach has been used to manage the relationships between the many different types of designs. Theorems of particular interest include the necessary conditions for existence of a BIBD, Wilson's construction on GDDs, and Bose's inequality on resolvable designs. Parts of this formalisation are explored in the paper "A Modular First Formalisation of Combinatorial Design Theory", presented at CICM 2021. Relational Forests https://www.isa-afp.org/entries/Relational_Forests.html https://www.isa-afp.org/entries/Relational_Forests.html Walter Guttmann 03 Aug 2021 00:00:00 +0000 We study second-order formalisations of graph properties expressed as first-order formulas in relation algebras extended with a Kleene star. The formulas quantify over relations while still avoiding quantification over elements of the base set. We formalise the property of undirected graphs being acyclic this way. This involves a study of various kinds of orientation of graphs. We also verify basic algorithms to constructively prove several second-order properties. - - Schutz' Independent Axioms for Minkowski Spacetime - https://www.isa-afp.org/entries/Schutz_Spacetime.html - https://www.isa-afp.org/entries/Schutz_Spacetime.html - Richard Schmoetten, Jake Palmer, Jacques Fleuriot - 27 Jul 2021 00:00:00 +0000 - -This is a formalisation of Schutz' system of axioms for Minkowski -spacetime published under the name "Independent axioms for -Minkowski space-time" in 1997, as well as most of the results in -the third chapter ("Temporal Order on a Path") of the above -monograph. Many results are proven here that cannot be found in -Schutz, either preceding the theorem they are needed for, or within -their own thematic section. - - - Finitely Generated Abelian Groups - https://www.isa-afp.org/entries/Finitely_Generated_Abelian_Groups.html - https://www.isa-afp.org/entries/Finitely_Generated_Abelian_Groups.html - Joseph Thommes, Manuel Eberl - 07 Jul 2021 00:00:00 +0000 - -This article deals with the formalisation of some group-theoretic -results including the fundamental theorem of finitely generated -abelian groups characterising the structure of these groups as a -uniquely determined product of cyclic groups. Both the invariant -factor decomposition and the primary decomposition are covered. -Additional work includes results about the direct product, the -internal direct product and more group-theoretic lemmas. - - - SpecCheck - Specification-Based Testing for Isabelle/ML - https://www.isa-afp.org/entries/SpecCheck.html - https://www.isa-afp.org/entries/SpecCheck.html - Kevin Kappelmann, Lukas Bulwahn, Sebastian Willenbrink - 01 Jul 2021 00:00:00 +0000 - -SpecCheck is a <a -href="https://en.wikipedia.org/wiki/QuickCheck">QuickCheck</a>-like -testing framework for Isabelle/ML. You can use it to write -specifications for ML functions. SpecCheck then checks whether your -specification holds by testing your function against a given number of -generated inputs. It helps you to identify bugs by printing -counterexamples on failure and provides you timing information. -SpecCheck is customisable and allows you to specify your own input -generators, test output formats, as well as pretty printers and -shrinking functions for counterexamples among other things. - - - Van der Waerden's Theorem - https://www.isa-afp.org/entries/Van_der_Waerden.html - https://www.isa-afp.org/entries/Van_der_Waerden.html - Katharina Kreuzer, Manuel Eberl - 22 Jun 2021 00:00:00 +0000 - -This article formalises the proof of Van der Waerden's Theorem -from Ramsey theory. Van der Waerden's Theorem states that for -integers $k$ and $l$ there exists a number $N$ which guarantees that -if an integer interval of length at least $N$ is coloured with $k$ -colours, there will always be an arithmetic progression of length $l$ -of the same colour in said interval. The proof goes along the lines of -\cite{Swan}. The smallest number $N_{k,l}$ fulfilling Van der -Waerden's Theorem is then called the Van der Waerden Number. -Finding the Van der Waerden Number is still an open problem for most -values of $k$ and $l$. - - - MiniSail - A kernel language for the ISA specification language SAIL - https://www.isa-afp.org/entries/MiniSail.html - https://www.isa-afp.org/entries/MiniSail.html - Mark Wassell - 18 Jun 2021 00:00:00 +0000 - -MiniSail is a kernel language for Sail, an instruction set -architecture (ISA) specification language. Sail is an imperative -language with a light-weight dependent type system similar to -refinement type systems. From an ISA specification, the Sail compiler -can generate theorem prover code and C (or OCaml) to give an -executable emulator for an architecture. The idea behind MiniSail is -to capture the key and novel features of Sail in terms of their -syntax, typing rules and operational semantics, and to confirm that -they work together by proving progress and preservation lemmas. We use -the Nominal2 library to handle binding. - - - Public Announcement Logic - https://www.isa-afp.org/entries/Public_Announcement_Logic.html - https://www.isa-afp.org/entries/Public_Announcement_Logic.html - Asta Halkjær From - 17 Jun 2021 00:00:00 +0000 - -This work is a formalization of public announcement logic with -countably many agents. It includes proofs of soundness and -completeness for a variant of the axiom system PA + DIST! + NEC!. The -completeness proof builds on the Epistemic Logic theory. - - - A Shorter Compiler Correctness Proof for Language IMP - https://www.isa-afp.org/entries/IMP_Compiler.html - https://www.isa-afp.org/entries/IMP_Compiler.html - Pasquale Noce - 04 Jun 2021 00:00:00 +0000 - -This paper presents a compiler correctness proof for the didactic -imperative programming language IMP, introduced in Nipkow and -Klein's book on formal programming language semantics (version of -March 2021), whose size is just two thirds of the book's proof in -the number of formal text lines. As such, it promises to constitute a -further enhanced reference for the formal verification of compilers -meant for larger, real-world programming languages. The presented -proof does not depend on language determinism, so that the proposed -approach can be applied to non-deterministic languages as well. As a -confirmation, this paper extends IMP with an additional -non-deterministic choice command, and proves compiler correctness, -viz. the simulation of compiled code execution by source code, for -such extended language. - - - Lyndon words - https://www.isa-afp.org/entries/Combinatorics_Words_Lyndon.html - https://www.isa-afp.org/entries/Combinatorics_Words_Lyndon.html - Štěpán Holub, Štěpán Starosta - 24 May 2021 00:00:00 +0000 - -Lyndon words are words lexicographically minimal in their conjugacy -class. We formalize their basic properties and characterizations, in -particular the concepts of the longest Lyndon suffix and the Lyndon -factorization. Most of the work assumes a fixed lexicographical order. -Nevertheless we also define the smallest relation guaranteeing -lexicographical minimality of a given word (in its conjugacy class). - diff --git a/web/statistics.html b/web/statistics.html --- a/web/statistics.html +++ b/web/statistics.html @@ -1,307 +1,307 @@ Archive of Formal Proofs

 

 

 

 

 

 

Statistics

 

Statistics

- - - - + + + +
Number of Articles:631
Number of Authors:405
Number of lemmas:~184,700
Lines of Code:~3,184,100
Number of Articles:639
Number of Authors:409
Number of lemmas:~187,300
Lines of Code:~3,242,600

Most used AFP articles:

- + - + + + + + - + - - - -
NameUsed by ? articles
1. List-Index1819
2. Show1314
3. Coinductive 12
Collections 12
Regular-Sets 12
4. Jordan_Normal_Form 11
Landau_Symbols 11
Polynomial_Factorization11
5.Polynomial_FactorizationAbstract-Rewriting 10
6.Abstract-Rewriting9
Automatic_Refinement 9
Deriving 9
Native_Word 9

Growth in number of articles:

Growth in lines of code:

Growth in number of authors:

Size of articles:

\ No newline at end of file diff --git a/web/topics.html b/web/topics.html --- a/web/topics.html +++ b/web/topics.html @@ -1,1002 +1,1007 @@ Archive of Formal Proofs

 

 

 

 

 

 

Index by Topic

 

Computer science

Artificial intelligence

Automata and formal languages

Algorithms

Knuth_Morris_Pratt   Probabilistic_While   Comparison_Sort_Lower_Bound   Quick_Sort_Cost   TortoiseHare   Selection_Heap_Sort   VerifyThis2018   CYK   Boolean_Expression_Checkers   Efficient-Mergesort   SATSolverVerification   MuchAdoAboutTwo   First_Order_Terms   Monad_Memo_DP   Hidden_Markov_Models   Imperative_Insertion_Sort   Formal_SSA   ROBDD   Median_Of_Medians_Selection   Fisher_Yates   Optimal_BST   IMP2   Auto2_Imperative_HOL   List_Inversions   IMP2_Binary_Heap   MFOTL_Monitor   Adaptive_State_Counting   Generic_Join   VerifyThis2019   Generalized_Counting_Sort   MFODL_Monitor_Optimized   Sliding_Window_Algorithm   PAC_Checker   Regression_Test_Selection   Graph: DFS_Framework   Prpu_Maxflow   Floyd_Warshall   Roy_Floyd_Warshall   Dijkstra_Shortest_Path   EdmondsKarp_Maxflow   Depth-First-Search   GraphMarkingIBP   Transitive-Closure   Transitive-Closure-II   Gabow_SCC   Kruskal   Prim_Dijkstra_Simple   Relational_Minimum_Spanning_Trees   Distributed: DiskPaxos   GenClock   ClockSynchInst   Heard_Of   Consensus_Refined   Abortable_Linearizable_Modules   IMAP-CRDT   CRDT   Chandy_Lamport   OpSets   Stellar_Quorums   WOOT_Strong_Eventual_Consistency   Progress_Tracking   Concurrent: ConcurrentGC   Online: List_Update   Geometry: Closest_Pair_Points   Approximation: Approximation_Algorithms   Mathematical: FFT   Gauss-Jordan-Elim-Fun   UpDown_Scheme   Polynomials   Gauss_Jordan   Echelon_Form   QR_Decomposition   Hermite   Groebner_Bases   Diophantine_Eqns_Lin_Hom   Taylor_Models   LLL_Basis_Reduction   Signature_Groebner   BenOr_Kozen_Reif   Smith_Normal_Form   Safe_Distance   Modular_arithmetic_LLL_and_HNF_algorithms   Virtual_Substitution   Optimization: Simplex   Quantum computing: Isabelle_Marries_Dirac   Projective_Measurements   Registers  

Concurrency

Data structures

Functional programming

Hardware

Machine learning

Networks

Programming languages

Clean   Decl_Sem_Fun_PL   Language definitions: CakeML   WebAssembly   pGCL   GPU_Kernel_PL   LightweightJava   CoreC++   FeatherweightJava   Jinja   JinjaThreads   Locally-Nameless-Sigma   AutoFocus-Stream   FocusStreamsCaseStudies   Isabelle_Meta_Model   Simpl   Complx   Safe_OCL   Isabelle_C   JinjaDCI   Lambda calculi: Higher_Order_Terms   Launchbury   PCF   POPLmark-deBruijn   Lam-ml-Normalization   LambdaMu   Binding_Syntax_Theory   LambdaAuth   Type systems: Name_Carrying_Type_Inference   MiniML   Possibilistic_Noninterference   SIFUM_Type_Systems   Dependent_SIFUM_Type_Systems   Strong_Security   WHATandWHERE_Security   VolpanoSmith   Physical_Quantities   MiniSail   Logics: ConcurrentIMP   Refine_Monadic   Automatic_Refinement   MonoBoolTranAlgebra   Simpl   Separation_Algebra   Separation_Logic_Imperative_HOL   Relational-Incorrectness-Logic   Abstract-Hoare-Logics   Kleene_Algebra   KAT_and_DRA   KAD   BytecodeLogicJmlTypes   DataRefinementIBP   RefinementReactive   SIFPL   TLA   Ribbon_Proofs   Separata   Complx   Differential_Dynamic_Logic   Hoare_Time   IMP2   UTP   QHLProver   Differential_Game_Logic   Correctness_Algebras   Registers   Compiling: CakeML_Codegen   Compiling-Exceptions-Correctly   NormByEval   Density_Compiler   VeriComp   IMP_Compiler   Static analysis: RIPEMD-160-SPARK   Program-Conflict-Analysis   Shivers-CFA   Slicing   HRB-Slicing   InfPathElimination   Abs_Int_ITP2012   Dominance_CHK   Transformations: Call_Arity   Refine_Imperative_HOL   WorkerWrapper   Monad_Memo_DP   Formal_SSA   Minimal_SSA   Misc: JiveDataStoreModel   Pop_Refinement   Case_Labeling   Interpreter_Optimizations  

Security

Semantics

System description languages

Logic

Philosophical aspects

General logic

Computability

Set theory

Proof theory

Rewriting

Mathematics

Order

Algebra

Analysis

Probability theory

Number theory

Games and economics

Geometry

Topology

Graph theory

Combinatorics

Category theory

Physics

Misc

Tools

\ No newline at end of file